From 7cfc4e5146be5666419451bdd516f1f3f264d24a Mon Sep 17 00:00:00 2001 From: Enrico Tassi Date: Sun, 25 Jan 2015 14:42:51 +0100 Subject: Imported Upstream version 8.5~beta1+dfsg --- library/assumptions.ml | 134 ++-- library/assumptions.mli | 18 +- library/decl_kinds.ml | 125 ---- library/decl_kinds.mli | 90 --- library/declare.ml | 356 ++++++--- library/declare.mli | 63 +- library/declaremods.ml | 1594 +++++++++++++++++++---------------------- library/declaremods.mli | 152 ++-- library/decls.ml | 38 +- library/decls.mli | 12 +- library/dischargedhypsmap.ml | 34 +- library/dischargedhypsmap.mli | 9 +- library/global.ml | 294 +++++--- library/global.mli | 165 +++-- library/globnames.ml | 247 +++++++ library/globnames.mli | 103 +++ library/goptions.ml | 143 ++-- library/goptions.mli | 26 +- library/goptionstyp.mli | 26 - library/heads.ml | 85 +-- library/heads.mli | 2 +- library/impargs.ml | 190 ++--- library/impargs.mli | 18 +- library/keys.ml | 170 +++++ library/keys.mli | 23 + library/kindops.ml | 67 ++ library/kindops.mli | 15 + library/lib.ml | 437 +++++------ library/lib.mli | 127 ++-- library/libnames.ml | 315 +++----- library/libnames.mli | 159 ++-- library/libobject.ml | 54 +- library/libobject.mli | 8 +- library/library.ml | 646 ++++++++++------- library/library.mli | 61 +- library/library.mllib | 8 +- library/loadpath.ml | 135 ++++ library/loadpath.mli | 58 ++ library/nameops.ml | 51 +- library/nameops.mli | 47 +- library/nametab.ml | 351 ++++----- library/nametab.mli | 49 +- library/states.ml | 36 +- library/states.mli | 19 +- library/summary.ml | 171 ++++- library/summary.mli | 59 +- library/universes.ml | 1006 ++++++++++++++++++++++++++ library/universes.mli | 253 +++++++ 48 files changed, 5131 insertions(+), 3118 deletions(-) delete mode 100644 library/decl_kinds.ml delete mode 100644 library/decl_kinds.mli create mode 100644 library/globnames.ml create mode 100644 library/globnames.mli delete mode 100644 library/goptionstyp.mli create mode 100644 library/keys.ml create mode 100644 library/keys.mli create mode 100644 library/kindops.ml create mode 100644 library/kindops.mli create mode 100644 library/loadpath.ml create mode 100644 library/loadpath.mli create mode 100644 library/universes.ml create mode 100644 library/universes.mli (limited to 'library') diff --git a/library/assumptions.ml b/library/assumptions.ml index 273ddb55..04ee14fb 100644 --- a/library/assumptions.ml +++ b/library/assumptions.ml @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* id_ord i1 i2 - | Axiom k1 , Axiom k2 -> cst_ord k1 k2 - | Opaque k1 , Opaque k2 -> cst_ord k1 k2 - | Variable _ , Axiom _ -> -1 + | Variable i1 , Variable i2 -> Id.compare i1 i2 + | Axiom k1 , Axiom k2 -> con_ord k1 k2 + | Opaque k1 , Opaque k2 -> con_ord k1 k2 + | Transparent k1 , Transparent k2 -> con_ord k1 k2 | Axiom _ , Variable _ -> 1 - | Opaque _ , _ -> -1 - | _, Opaque _ -> 1 + | Opaque _ , Variable _ + | Opaque _ , Axiom _ -> 1 + | Transparent _ , Variable _ + | Transparent _ , Axiom _ + | Transparent _ , Opaque _ -> 1 + | _ , _ -> -1 end module ContextObjectSet = Set.Make (OrderedContextObject) @@ -56,14 +59,31 @@ let modcache = ref (MPmap.empty : structure_body MPmap.t) let rec search_mod_label lab = function | [] -> raise Not_found - | (l,SFBmodule mb) :: _ when l = lab -> mb + | (l, SFBmodule mb) :: _ when Label.equal l lab -> mb | _ :: fields -> search_mod_label lab fields let rec search_cst_label lab = function | [] -> raise Not_found - | (l,SFBconst cb) :: _ when l = lab -> cb + | (l, SFBconst cb) :: _ when Label.equal l lab -> cb | _ :: fields -> search_cst_label lab fields +(* TODO: using [empty_delta_resolver] below is probably slightly incorrect. But: + a) I don't see currently what should be used instead + b) this shouldn't be critical for Print Assumption. At worse some + constants will have a canonical name which is non-canonical, + leading to failures in [Global.lookup_constant], but our own + [lookup_constant] should work. +*) + +let rec fields_of_functor f subs mp0 args = function + |NoFunctor a -> f subs mp0 args a + |MoreFunctor (mbid,_,e) -> + match args with + | [] -> assert false (* we should only encounter applied functors *) + | mpa :: args -> + let subs = add_mbid mbid mpa empty_delta_resolver (*TODO*) subs in + fields_of_functor f subs mp0 args e + let rec lookup_module_in_impl mp = try Global.lookup_module mp with Not_found -> @@ -87,46 +107,32 @@ and fields_of_mp mp = let mb = lookup_module_in_impl mp in let fields,inner_mp,subs = fields_of_mb empty_subst mb [] in let subs = - if inner_mp = mp then subs + if mp_eq inner_mp mp then subs else add_mp inner_mp mp mb.mod_delta subs in - Modops.subst_signature subs fields + Modops.subst_structure subs fields -and fields_of_mb subs mb args = - let seb = match mb.mod_expr with - | None -> mb.mod_type (* cf. Declare Module *) - | Some seb -> seb - in - fields_of_seb subs mb.mod_mp seb args +and fields_of_mb subs mb args = match mb.mod_expr with + |Algebraic expr -> fields_of_expression subs mb.mod_mp args expr + |Struct sign -> fields_of_signature subs mb.mod_mp args sign + |Abstract|FullStruct -> fields_of_signature subs mb.mod_mp args mb.mod_type -(* TODO: using [empty_delta_resolver] below in [fields_of_seb] - is probably slightly incorrect. But: - a) I don't see currently what should be used instead - b) this shouldn't be critical for Print Assumption. At worse some - constants will have a canonical name which is non-canonical, - leading to failures in [Global.lookup_constant], but our own - [lookup_constant] should work. -*) +(** The Abstract case above corresponds to [Declare Module] *) -and fields_of_seb subs mp0 seb args = match seb with - | SEBstruct l -> - assert (args = []); - l, mp0, subs - | SEBident mp -> +and fields_of_signature x = + fields_of_functor + (fun subs mp0 args struc -> + assert (List.is_empty args); + (struc, mp0, subs)) x + +and fields_of_expr subs mp0 args = function + |MEident mp -> let mb = lookup_module_in_impl (subst_mp subs mp) in fields_of_mb subs mb args - | SEBapply (seb1,seb2,_) -> - (match seb2 with - | SEBident mp2 -> fields_of_seb subs mp0 seb1 (mp2::args) - | _ -> assert false) (* only legal application is to module names *) - | SEBfunctor (mbid,mtb,seb) -> - (match args with - | [] -> assert false (* we should only encounter applied functors *) - | mpa :: args -> - let subs = add_mbid mbid mpa empty_delta_resolver subs in - fields_of_seb subs mp0 seb args) - | SEBwith _ -> assert false (* should not appear in a mod_expr - or mod_type field *) + |MEapply (me1,mp2) -> fields_of_expr subs mp0 (mp2::args) me1 + |MEwith _ -> assert false (* no 'with' in [mod_expr] *) + +and fields_of_expression x = fields_of_functor fields_of_expr x let lookup_constant_in_impl cst fallback = try @@ -143,16 +149,16 @@ let lookup_constant_in_impl cst fallback = - The label has not been found in the structure. This is an error *) match fallback with | Some cb -> cb - | None -> anomaly ("Print Assumption: unknown constant "^string_of_con cst) + | None -> anomaly (str "Print Assumption: unknown constant " ++ pr_con cst) let lookup_constant cst = try let cb = Global.lookup_constant cst in - if constant_has_body cb then cb + if Declareops.constant_has_body cb then cb else lookup_constant_in_impl cst (Some cb) with Not_found -> lookup_constant_in_impl cst None -let assumptions ?(add_opaque=false) st (* t *) = +let assumptions ?(add_opaque=false) ?(add_transparent=false) st (* t *) = modcache := MPmap.empty; let (idts,knst) = st in (* Infix definition for chaining function that accumulate @@ -181,7 +187,7 @@ let assumptions ?(add_opaque=false) st (* t *) = a "Let" definition, in the former it is an assumption of [t], in the latter is must be unfolded like a Const. The other cases are straightforward recursion. - Calls to the environment are memoized, thus avoiding to explore + Calls to the environment are memoized, thus avoiding exploration of the DAG of the environment as if it was a tree (can cause exponential behavior and prevent the algorithm from terminating in reasonable time). [s] is a set of [context_object], representing @@ -198,7 +204,7 @@ let assumptions ?(add_opaque=false) st (* t *) = | Case (_,e1,e2,e_array) -> (iter e1)**(iter e2)**(iter_array e_array) | Fix (_,(_, e1_array, e2_array)) | CoFix (_,(_,e1_array, e2_array)) -> (iter_array e1_array) ** (iter_array e2_array) - | Const kn -> do_memoize_kn kn + | Const (kn,_) -> do_memoize_kn kn | _ -> identity2 (* closed atomic types + rel *) and iter_array a = Array.fold_right (fun e f -> (iter e)**f) a identity2 in iter t s acc @@ -216,23 +222,23 @@ let assumptions ?(add_opaque=false) st (* t *) = and add_kn kn s acc = let cb = lookup_constant kn in let do_type cst = - let ctype = - match cb.Declarations.const_type with - | PolymorphicArity (ctx,a) -> mkArity (ctx, Type a.poly_level) - | NonPolymorphicType t -> t - in + let ctype = Global.type_of_global_unsafe (Globnames.ConstRef kn) in (s,ContextObjectMap.add cst ctype acc) in let (s,acc) = - if add_opaque && Declarations.constant_has_body cb - && (Declarations.is_opaque cb || not (Cpred.mem kn knst)) - then - do_type (Opaque kn) - else (s,acc) + if Declareops.constant_has_body cb then + if Declareops.is_opaque cb || not (Cpred.mem kn knst) then + (** it is opaque *) + if add_opaque then do_type (Opaque kn) + else (s, acc) + else + if add_transparent then do_type (Transparent kn) + else (s, acc) + else (s, acc) in - match Declarations.body_of_constant cb with + match Global.body_of_constant_body cb with | None -> do_type (Axiom kn) - | Some body -> do_constr (Declarations.force body) s acc + | Some body -> do_constr body s acc and do_memoize_kn kn = try_and_go (Axiom kn) (add_kn kn) diff --git a/library/assumptions.mli b/library/assumptions.mli index 7756c575..0a2c62f5 100644 --- a/library/assumptions.mli +++ b/library/assumptions.mli @@ -1,28 +1,30 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* transparent_state -> constr -> + ?add_opaque:bool -> ?add_transparent:bool -> transparent_state -> constr -> Term.types ContextObjectMap.t diff --git a/library/decl_kinds.ml b/library/decl_kinds.ml deleted file mode 100644 index ee7acec4..00000000 --- a/library/decl_kinds.ml +++ /dev/null @@ -1,125 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* IsDefinition d - | Proof s -> IsProof s - -let string_of_theorem_kind = function - | Theorem -> "Theorem" - | Lemma -> "Lemma" - | Fact -> "Fact" - | Remark -> "Remark" - | Property -> "Property" - | Proposition -> "Proposition" - | Corollary -> "Corollary" - -let string_of_definition_kind def = - match def with - | Local, Coercion -> "Coercion Local" - | Global, Coercion -> "Coercion" - | Local, Definition -> "Let" - | Global, Definition -> "Definition" - | Local, SubClass -> "Local SubClass" - | Global, SubClass -> "SubClass" - | Global, CanonicalStructure -> "Canonical Structure" - | Global, Example -> "Example" - | Local, (CanonicalStructure|Example) -> - anomaly "Unsupported local definition kind" - | Local, Instance -> "Instance" - | Global, Instance -> "Global Instance" - | _, (StructureComponent|Scheme|CoFixpoint|Fixpoint|IdentityCoercion|Method) - -> anomaly "Internal definition kind" - -(* Strength *) - -let strength_of_global = function - | VarRef _ -> Local - | IndRef _ | ConstructRef _ | ConstRef _ -> Global - -let string_of_strength = function - | Local -> "Local" - | Global -> "Global" - - -(* Recursive power *) - -(* spiwack: this definition might be of use in the kernel, for now I do not - push them deeper than needed, though. *) -type recursivity_kind = - | Finite (* = inductive *) - | CoFinite (* = coinductive *) - | BiFinite (* = non-recursive, like in "Record" definitions *) - -(* helper, converts to "finiteness flag" booleans *) -let recursivity_flag_of_kind = function - | Finite | BiFinite -> true - | CoFinite -> false diff --git a/library/decl_kinds.mli b/library/decl_kinds.mli deleted file mode 100644 index 44f5cbab..00000000 --- a/library/decl_kinds.mli +++ /dev/null @@ -1,90 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* logical_kind -val string_of_theorem_kind : theorem_kind -> string -val string_of_definition_kind : - locality * definition_object_kind -> string - -(** About locality *) - -val strength_of_global : global_reference -> locality -val string_of_strength : locality -> string - -(** About recursive power of type declarations *) - -type recursivity_kind = - | Finite (** = inductive *) - | CoFinite (** = coinductive *) - | BiFinite (** = non-recursive, like in "Record" definitions *) - -(** helper, converts to "finiteness flag" booleans *) -val recursivity_flag_of_kind : recursivity_kind -> bool diff --git a/library/declare.ml b/library/declare.ml index d6413d3d..7f42a747 100644 --- a/library/declare.ml +++ b/library/declare.ml @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* ()) -let xml_declare_constant = ref (fun (sp:internal_flag * constant)-> ()) -let xml_declare_inductive = ref (fun (sp:internal_flag * object_name) -> ()) - -let if_xml f x = if !Flags.xml_export then f x else () - -let set_xml_declare_variable f = xml_declare_variable := if_xml f -let set_xml_declare_constant f = xml_declare_constant := if_xml f -let set_xml_declare_inductive f = xml_declare_inductive := if_xml f - -let cache_hook = ref ignore -let add_cache_hook f = cache_hook := f - (** Declaration of section variables and local definitions *) type section_variable_entry = - | SectionLocalDef of constr * types option * bool (* opacity *) - | SectionLocalAssum of types * bool (* Implicit status *) + | SectionLocalDef of definition_entry + | SectionLocalAssum of types Univ.in_universe_context_set * polymorphic * bool (** Implicit status *) -type variable_declaration = dir_path * section_variable_entry * logical_kind +type variable_declaration = DirPath.t * section_variable_entry * logical_kind let cache_variable ((sp,_),o) = match o with - | Inl cst -> Global.add_constraints cst + | Inl ctx -> Global.push_context_set ctx | Inr (id,(p,d,mk)) -> (* Constr raisonne sur les noms courts *) if variable_exists id then alreadydeclared (pr_id id ++ str " already exists"); - let impl,opaq,cst = match d with (* Fails if not well-typed *) - | SectionLocalAssum (ty, impl) -> - let cst = Global.push_named_assum (id,ty) in - let impl = if impl then Lib.Implicit else Lib.Explicit in - impl, true, cst - | SectionLocalDef (c,t,opaq) -> - let cst = Global.push_named_def (id,c,t) in - Lib.Explicit, opaq, cst in + + let impl,opaq,poly,ctx = match d with (* Fails if not well-typed *) + | SectionLocalAssum ((ty,ctx),poly,impl) -> + let () = Global.push_named_assum ((id,ty),ctx) in + let impl = if impl then Implicit else Explicit in + impl, true, poly, ctx + | SectionLocalDef (de) -> + let () = Global.push_named_def (id,de) in + Explicit, de.const_entry_opaque, de.const_entry_polymorphic, + (Univ.ContextSet.of_context de.const_entry_universes) in Nametab.push (Nametab.Until 1) (restrict_path 0 sp) (VarRef id); - add_section_variable id impl; + add_section_variable id impl poly ctx; Dischargedhypsmap.set_discharged_hyps sp []; - add_variable_data id (p,opaq,cst,mk) + add_variable_data id (p,opaq,ctx,poly,mk) let discharge_variable (_,o) = match o with - | Inr (id,_) -> Some (Inl (variable_constraints id)) + | Inr (id,_) -> + if variable_polymorphic id then None + else Some (Inl (variable_context id)) | Inl _ -> Some o type variable_obj = - (Univ.constraints, identifier * variable_declaration) union + (Univ.ContextSet.t, Id.t * variable_declaration) union let inVariable : variable_obj -> obj = declare_object { (default_object "VARIABLE") with @@ -93,70 +83,90 @@ let declare_variable id obj = declare_var_implicits id; Notation.declare_ref_arguments_scope (VarRef id); Heads.declare_head (EvalVarRef id); - !xml_declare_variable oname; oname (** Declaration of constants and parameters *) +type constant_obj = { + cst_decl : global_declaration; + cst_hyps : Dischargedhypsmap.discharged_hyps; + cst_kind : logical_kind; + cst_locl : bool; +} + type constant_declaration = constant_entry * logical_kind (* At load-time, the segment starting from the module name to the discharge *) (* section (if Remark or Fact) is needed to access a construction *) -let load_constant i ((sp,kn),(_,_,kind)) = +let load_constant i ((sp,kn), obj) = if Nametab.exists_cci sp then alreadydeclared (pr_id (basename sp) ++ str " already exists"); let con = Global.constant_of_delta_kn kn in Nametab.push (Nametab.Until i) sp (ConstRef con); - add_constant_kind con kind + add_constant_kind con obj.cst_kind (* Opening means making the name without its module qualification available *) -let open_constant i ((sp,kn),_) = - let con = Global.constant_of_delta_kn kn in - Nametab.push (Nametab.Exactly i) sp (ConstRef con) +let open_constant i ((sp,kn), obj) = + (** Never open a local definition *) + if obj.cst_locl then () + else + let con = Global.constant_of_delta_kn kn in + Nametab.push (Nametab.Exactly i) sp (ConstRef con); + match (Global.lookup_constant con).const_body with + | (Def _ | Undef _) -> () + | OpaqueDef lc -> + match Opaqueproof.get_constraints (Global.opaque_tables ())lc with + | Some f when Future.is_val f -> Global.push_context_set (Future.force f) + | _ -> () let exists_name id = - variable_exists id or Global.exists_objlabel (label_of_id id) + variable_exists id || Global.exists_objlabel (Label.of_id id) let check_exists sp = let id = basename sp in if exists_name id then alreadydeclared (pr_id id ++ str " already exists") -let cache_constant ((sp,kn),(cdt,dhyps,kind)) = +let cache_constant ((sp,kn), obj) = let id = basename sp in let _,dir,_ = repr_kn kn in - check_exists sp; - let kn' = Global.add_constant dir id cdt in - assert (kn' = constant_of_kn kn); + let () = check_exists sp in + let kn' = Global.add_constant dir id obj.cst_decl in + assert (eq_constant kn' (constant_of_kn kn)); Nametab.push (Nametab.Until 1) sp (ConstRef (constant_of_kn kn)); - add_section_constant kn' (Global.lookup_constant kn').const_hyps; - Dischargedhypsmap.set_discharged_hyps sp dhyps; - add_constant_kind (constant_of_kn kn) kind; - !cache_hook sp + let cst = Global.lookup_constant kn' in + add_section_constant (cst.const_proj <> None) kn' cst.const_hyps; + Dischargedhypsmap.set_discharged_hyps sp obj.cst_hyps; + add_constant_kind (constant_of_kn kn) obj.cst_kind let discharged_hyps kn sechyps = let (_,dir,_) = repr_kn kn in let args = Array.to_list (instance_from_variable_context sechyps) in - List.rev (List.map (Libnames.make_path dir) args) + List.rev_map (Libnames.make_path dir) args -let discharge_constant ((sp,kn),(cdt,dhyps,kind)) = +let discharge_constant ((sp, kn), obj) = let con = constant_of_kn kn in - let cb = Global.lookup_constant con in - let repl = replacement_context () in - let sechyps = section_segment_of_constant con in - let recipe = { d_from=cb; d_modlist=repl; d_abstract=named_of_variable_context sechyps } in - Some (GlobalRecipe recipe,(discharged_hyps kn sechyps)@dhyps,kind) + let from = Global.lookup_constant con in + let modlist = replacement_context () in + let hyps,subst,uctx = section_segment_of_constant con in + let new_hyps = (discharged_hyps kn hyps) @ obj.cst_hyps in + let abstract = (named_of_variable_context hyps, subst, uctx) in + let new_decl = GlobalRecipe{ from; info = { Opaqueproof.modlist; abstract}} in + Some { obj with cst_hyps = new_hyps; cst_decl = new_decl; } (* Hack to reduce the size of .vo: we keep only what load/open needs *) -let dummy_constant_entry = ConstantEntry (ParameterEntry (None,mkProp,None)) - -let dummy_constant (ce,_,mk) = dummy_constant_entry,[],mk +let dummy_constant_entry = + ConstantEntry (ParameterEntry (None,false,(mkProp,Univ.UContext.empty),None)) + +let dummy_constant cst = { + cst_decl = dummy_constant_entry; + cst_hyps = []; + cst_kind = cst.cst_kind; + cst_locl = cst.cst_locl; +} let classify_constant cst = Substitute (dummy_constant cst) -type constant_obj = - global_declaration * Dischargedhypsmap.discharged_hyps * logical_kind - let inConstant : constant_obj -> obj = declare_object { (default_object "CONSTANT") with cache_function = cache_constant; @@ -166,23 +176,125 @@ let inConstant : constant_obj -> obj = subst_function = ident_subst_function; discharge_function = discharge_constant } -let declare_constant_common id dhyps (cd,kind) = - let (sp,kn) = add_leaf id (inConstant (cd,dhyps,kind)) in +let declare_constant_common id cst = + let (sp,kn) = add_leaf id (inConstant cst) in let c = Global.constant_of_delta_kn kn in declare_constant_implicits c; Heads.declare_head (EvalConstRef c); Notation.declare_ref_arguments_scope (ConstRef c); c -let declare_constant ?(internal = UserVerbose) id (cd,kind) = - let kn = declare_constant_common id [] (ConstantEntry cd,kind) in - !xml_declare_constant (internal,kn); +let definition_entry ?(opaque=false) ?(inline=false) ?types + ?(poly=false) ?(univs=Univ.UContext.empty) ?(eff=Declareops.no_seff) body = + { const_entry_body = Future.from_val ((body,Univ.ContextSet.empty), eff); + const_entry_secctx = None; + const_entry_type = types; + const_entry_polymorphic = poly; + const_entry_universes = univs; + const_entry_opaque = opaque; + const_entry_feedback = None; + const_entry_inline_code = inline} + +let declare_scheme = ref (fun _ _ -> assert false) +let set_declare_scheme f = declare_scheme := f +let declare_sideff env fix_exn se = + let cbl, scheme = match se with + | SEsubproof (c, cb, pt) -> [c, cb, pt], None + | SEscheme (cbl, k) -> + List.map (fun (_,c,cb,pt) -> c,cb,pt) cbl, Some (cbl,k) in + let id_of c = Names.Label.to_id (Names.Constant.label c) in + let pt_opaque_of cb pt = + match cb, pt with + | { const_body = Def sc }, _ -> (Mod_subst.force_constr sc, Univ.ContextSet.empty), false + | { const_body = OpaqueDef _ }, `Opaque(pt,univ) -> (pt, univ), true + | _ -> assert false + in + let ty_of cb = + match cb.Declarations.const_type with + | Declarations.RegularArity t -> Some t + | Declarations.TemplateArity _ -> None in + let cst_of cb pt = + let pt, opaque = pt_opaque_of cb pt in + let univs, subst = + if cb.const_polymorphic then + let univs = Univ.instantiate_univ_context cb.const_universes in + univs, Vars.subst_instance_constr (Univ.UContext.instance univs) + else cb.const_universes, fun x -> x + in + let pt = (subst (fst pt), snd pt) in + let ty = Option.map subst (ty_of cb) in + { cst_decl = ConstantEntry (DefinitionEntry { + const_entry_body = Future.from_here ~fix_exn (pt, Declareops.no_seff); + const_entry_secctx = Some cb.Declarations.const_hyps; + const_entry_type = ty; + const_entry_opaque = opaque; + const_entry_inline_code = false; + const_entry_feedback = None; + const_entry_polymorphic = cb.const_polymorphic; + const_entry_universes = univs; + }); + cst_hyps = [] ; + cst_kind = Decl_kinds.IsDefinition Decl_kinds.Definition; + cst_locl = true; + } in + let exists c = + try ignore(Environ.lookup_constant c env); true + with Not_found -> false in + let knl = + CList.map_filter (fun (c,cb,pt) -> + if exists c then None + else Some (c,declare_constant_common (id_of c) (cst_of cb pt))) cbl in + match scheme with + | None -> () + | Some (inds_consts,kind) -> + !declare_scheme kind (Array.of_list + (List.map (fun (c,kn) -> + CList.find_map (fun (x,c',_,_) -> + if Constant.equal c c' then Some (x,kn) else None) inds_consts) + knl)) + +let declare_constant ?(internal = UserVerbose) ?(local = false) id (cd, kind) = + let cd = (* We deal with side effects of non-opaque constants *) + match cd with + | Entries.DefinitionEntry ({ + const_entry_opaque = false; const_entry_body = bo } as de) + | Entries.DefinitionEntry ({ + const_entry_polymorphic = true; const_entry_body = bo } as de) + -> + let _, seff = Future.force bo in + if Declareops.side_effects_is_empty seff then cd + else begin + let seff = Declareops.uniquize_side_effects seff in + Declareops.iter_side_effects + (declare_sideff (Global.env ()) (Future.fix_exn_of bo)) seff; + Entries.DefinitionEntry { de with + const_entry_body = Future.chain ~pure:true bo (fun (pt, _) -> + pt, Declareops.no_seff) } + end + | _ -> cd + in + let cst = { + cst_decl = ConstantEntry cd; + cst_hyps = [] ; + cst_kind = kind; + cst_locl = local; + } in + let kn = declare_constant_common id cst in kn +let declare_definition ?(internal=UserVerbose) + ?(opaque=false) ?(kind=Decl_kinds.Definition) ?(local = false) + ?(poly=false) id ?types (body,ctx) = + let cb = + definition_entry ?types ~poly ~univs:(Univ.ContextSet.to_context ctx) ~opaque body + in + declare_constant ~internal ~local id + (Entries.DefinitionEntry cb, Decl_kinds.IsDefinition kind) + (** Declaration of inductive blocks *) let declare_inductive_argument_scopes kn mie = - list_iter_i (fun i {mind_entry_consnames=lc} -> + List.iteri (fun i {mind_entry_consnames=lc} -> Notation.declare_ref_arguments_scope (IndRef (kn,i)); for j=1 to List.length lc do Notation.declare_ref_arguments_scope (ConstructRef ((kn,i),j)); @@ -223,24 +335,24 @@ let cache_inductive ((sp,kn),(dhyps,mie)) = let id = basename sp in let _,dir,_ = repr_kn kn in let kn' = Global.add_mind dir id mie in - assert (kn'= mind_of_kn kn); - add_section_kn kn' (Global.lookup_mind kn').mind_hyps; + assert (eq_mind kn' (mind_of_kn kn)); + let mind = Global.lookup_mind kn' in + add_section_kn kn' mind.mind_hyps; Dischargedhypsmap.set_discharged_hyps sp dhyps; - List.iter (fun (sp, ref) -> Nametab.push (Nametab.Until 1) sp ref) names; - List.iter (fun (sp,_) -> !cache_hook sp) (inductive_names sp kn mie) - + List.iter (fun (sp, ref) -> Nametab.push (Nametab.Until 1) sp ref) names let discharge_inductive ((sp,kn),(dhyps,mie)) = let mind = Global.mind_of_delta_kn kn in let mie = Global.lookup_mind mind in let repl = replacement_context () in - let sechyps = section_segment_of_mutual_inductive mind in + let sechyps,usubst,uctx = section_segment_of_mutual_inductive mind in Some (discharged_hyps kn sechyps, - Discharge.process_inductive (named_of_variable_context sechyps) repl mie) + Discharge.process_inductive (named_of_variable_context sechyps,uctx) repl mie) let dummy_one_inductive_entry mie = { mind_entry_typename = mie.mind_entry_typename; mind_entry_arity = mkProp; + mind_entry_template = false; mind_entry_consnames = mie.mind_entry_consnames; mind_entry_lc = [] } @@ -248,9 +360,12 @@ let dummy_one_inductive_entry mie = { (* Hack to reduce the size of .vo: we keep only what load/open needs *) let dummy_inductive_entry (_,m) = ([],{ mind_entry_params = []; - mind_entry_record = false; - mind_entry_finite = true; - mind_entry_inds = List.map dummy_one_inductive_entry m.mind_entry_inds }) + 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_polymorphic = false; + mind_entry_universes = Univ.UContext.empty; + mind_entry_private = None }) type inductive_obj = Dischargedhypsmap.discharged_hyps * mutual_inductive_entry @@ -263,25 +378,38 @@ let inInductive : inductive_obj -> obj = subst_function = ident_subst_function; discharge_function = discharge_inductive } +let declare_projections mind = + let spec,_ = Inductive.lookup_mind_specif (Global.env ()) (mind,0) in + match spec.mind_record with + | Some (Some (_, kns, pjs)) -> + Array.iteri (fun i kn -> + let id = Label.to_id (Constant.label kn) in + let entry = {proj_entry_ind = mind; proj_entry_arg = i} in + let kn' = declare_constant id (ProjectionEntry entry, + IsDefinition StructureComponent) + in + assert(eq_constant kn kn')) kns; true + | Some None | None -> false + (* for initial declaration *) -let declare_mind isrecord mie = +let declare_mind mie = let id = match mie.mind_entry_inds with | ind::_ -> ind.mind_entry_typename - | [] -> anomaly "cannot declare an empty list of inductives" in + | [] -> anomaly (Pp.str "cannot declare an empty list of inductives") in let (sp,kn as oname) = add_leaf id (inInductive ([],mie)) in let mind = Global.mind_of_delta_kn kn in + let isprim = declare_projections mind in declare_mib_implicits mind; declare_inductive_argument_scopes mind mie; - !xml_declare_inductive (isrecord,oname); - oname + oname, isprim (* Declaration messages *) -let pr_rank i = str (ordinal (i+1)) +let pr_rank i = pr_nth (i+1) let fixpoint_message indexes l = - Flags.if_verbose msgnl (match l with - | [] -> anomaly "no recursive definition" + Flags.if_verbose msg_info (match l with + | [] -> anomaly (Pp.str "no recursive definition") | [id] -> pr_id id ++ str " is recursively defined" ++ (match indexes with | Some [|i|] -> str " (decreasing on "++pr_rank i++str " argument)" @@ -290,13 +418,13 @@ let fixpoint_message indexes l = spc () ++ str "are recursively defined" ++ match indexes with | Some a -> spc () ++ str "(decreasing respectively on " ++ - prlist_with_sep pr_comma pr_rank (Array.to_list a) ++ + prvect_with_sep pr_comma pr_rank a ++ str " arguments)" | None -> mt ())) let cofixpoint_message l = - Flags.if_verbose msgnl (match l with - | [] -> anomaly "No corecursive definition." + Flags.if_verbose msg_info (match l with + | [] -> anomaly (Pp.str "No corecursive definition.") | [id] -> pr_id id ++ str " is corecursively defined" | l -> hov 0 (prlist_with_sep pr_comma pr_id l ++ spc () ++ str "are corecursively defined")) @@ -305,7 +433,57 @@ let recursive_message isfix i l = (if isfix then fixpoint_message i else cofixpoint_message) l let definition_message id = - Flags.if_verbose msgnl (pr_id id ++ str " is defined") + Flags.if_verbose msg_info (pr_id id ++ str " is defined") let assumption_message id = - Flags.if_verbose msgnl (pr_id id ++ str " is assumed") + Flags.if_verbose msg_info (pr_id id ++ str " is assumed") + +(** Global universe names, in a different summary *) + +type universe_names = + (Univ.universe_level Idmap.t * Id.t Univ.LMap.t) + +let input_universes : universe_names -> Libobject.obj = + let open Libobject in + declare_object + { (default_object "Global universe name state") with + cache_function = (fun (na, pi) -> Universes.set_global_universe_names pi); + load_function = (fun _ (_, pi) -> Universes.set_global_universe_names pi); + discharge_function = (fun (_, a) -> Some a); + classify_function = (fun a -> Keep a) } + +let do_universe l = + let glob = Universes.global_universe_names () in + let glob' = + List.fold_left (fun (idl,lid) (l, id) -> + let lev = Universes.new_univ_level (Global.current_dirpath ()) in + (Idmap.add id lev idl, Univ.LMap.add lev id lid)) + glob l + in + Lib.add_anonymous_leaf (input_universes glob') + + +let input_constraints : Univ.constraints -> Libobject.obj = + let open Libobject in + declare_object + { (default_object "Global universe constraints") with + cache_function = (fun (na, c) -> Global.add_constraints c); + load_function = (fun _ (_, c) -> Global.add_constraints c); + discharge_function = (fun (_, a) -> Some a); + classify_function = (fun a -> Keep a) } + +let do_constraint l = + let u_of_id = + let names, _ = Universes.global_universe_names () in + fun (loc, id) -> + try Idmap.find id names + with Not_found -> + user_err_loc (loc, "Constraint", str "Undeclared universe " ++ pr_id id) + in + let constraints = List.fold_left (fun acc (l, d, r) -> + let lu = u_of_id l and ru = u_of_id r in + Univ.Constraint.add (lu, d, ru) acc) + Univ.Constraint.empty l + in + Lib.add_anonymous_leaf (input_constraints constraints) + diff --git a/library/declare.mli b/library/declare.mli index 2de128bd..03b66271 100644 --- a/library/declare.mli +++ b/library/declare.mli @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* variable_declaration -> object_name @@ -54,29 +47,43 @@ type internal_flag = | KernelSilent | UserVerbose +(* Defaut definition entries, transparent with no secctx or proj information *) +val definition_entry : ?opaque:bool -> ?inline:bool -> ?types:types -> + ?poly:polymorphic -> ?univs:Univ.universe_context -> ?eff:Declareops.side_effects -> + constr -> definition_entry + val declare_constant : - ?internal:internal_flag -> identifier -> constant_declaration -> constant + ?internal:internal_flag -> ?local:bool -> Id.t -> constant_declaration -> constant -(** [declare_mind me] declares a block of inductive types with - their constructors in the current section; it returns the path of - the whole block (boolean must be true iff it is a record) *) -val declare_mind : internal_flag -> mutual_inductive_entry -> object_name +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 -(** Hooks for XML output *) -val set_xml_declare_variable : (object_name -> unit) -> unit -val set_xml_declare_constant : (internal_flag * constant -> unit) -> unit -val set_xml_declare_inductive : (internal_flag * object_name -> unit) -> unit +(** Since transparent constant's side effects are globally declared, we + * need that *) +val set_declare_scheme : + (string -> (inductive * constant) array -> unit) -> unit -(** Hook for the cache function of constants and inductives *) -val add_cache_hook : (full_path -> unit) -> unit +(** [declare_mind me] declares a block of inductive types with + their constructors in the current section; it returns the path of + the whole block and a boolean indicating if it is a primitive record. *) +val declare_mind : mutual_inductive_entry -> object_name * bool (** Declaration messages *) -val definition_message : identifier -> unit -val assumption_message : identifier -> unit -val fixpoint_message : int array option -> identifier list -> unit -val cofixpoint_message : identifier list -> unit +val definition_message : Id.t -> unit +val assumption_message : Id.t -> unit +val fixpoint_message : int array option -> Id.t list -> unit +val cofixpoint_message : Id.t list -> unit val recursive_message : bool (** true = fixpoint *) -> - int array option -> identifier list -> unit + int array option -> Id.t list -> unit + +val exists_name : Id.t -> bool + + + +(** Global universe names and constraints *) -val exists_name : identifier -> bool +val do_universe : Id.t Loc.located list -> unit +val do_constraint : (Id.t Loc.located * Univ.constraint_type * Id.t Loc.located) list -> unit diff --git a/library/declaremods.ml b/library/declaremods.ml index 6b33e4b7..cc7c4d7f 100644 --- a/library/declaremods.ml +++ b/library/declaremods.ml @@ -1,53 +1,24 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* add_scope_subst sc1 sc2) scl - -let subst_scope sc = - try Stringmap.find sc !scope_subst with Not_found -> sc - -let reset_scope_subst () = - scope_subst := Stringmap.empty - -(** Which inline annotations should we honor, either None or the ones - whose level is less or equal to the given integer *) - -type inline = - | NoInline - | DefaultInline - | InlineAt of int +(** {6 Inlining levels} *) let default_inline () = Some (Flags.get_inline_level ()) @@ -56,57 +27,93 @@ let inl2intopt = function | InlineAt i -> Some i | DefaultInline -> default_inline () -type funct_app_annot = - { ann_inline : inline; - ann_scope_subst : scope_subst } +(** {6 Substitutive objects} -let inline_annot a = inl2intopt a.ann_inline + - The list of bound identifiers is nonempty only if the objects + are owned by a functor -type 'a annotated = ('a * funct_app_annot) + - Then comes either the object segment itself (for interactive + modules), or a compact way to store derived objects (path to + a earlier module + subtitution). +*) +type algebraic_objects = + | Objs of Lib.lib_objects + | Ref of module_path * substitution -(* modules and components *) +type substitutive_objects = MBId.t list * algebraic_objects -(* OBSOLETE This type is a functional closure of substitutive lib_objects. +(** ModSubstObjs : a cache of module substitutive objects - The first part is a partial substitution (which will be later - applied to lib_objects when completed). + This table is common to modules and module types. + - For a Module M:=N, the objects of N will be reloaded + with M after substitution. + - For a Module M:SIG:=..., the module M gets its objects from SIG - The second one is a list of bound identifiers which is nonempty - only if the objects are owned by a fuctor + Invariants: + - A alias (i.e. a module path inside a Ref constructor) should + never lead to another alias, but rather to a concrete Objs + constructor. - The third one is the "self" ident of the signature (or structure), - which should be substituted in lib_objects with the real name of - the module. + We will plug later a handler dealing with missing entries in the + cache. Such missing entries may come from inner parts of module + types, which aren't registered by the standard libobject machinery. +*) - The fourth one is the segment itself which can contain references - to identifiers in the domain of the substitution or in other two - parts. These references are invalid in the current scope and - therefore must be substitued with valid names before use. +module ModSubstObjs : + sig + val set : module_path -> substitutive_objects -> unit + val get : module_path -> substitutive_objects + val set_missing_handler : (module_path -> substitutive_objects) -> unit + end = + struct + let table = + Summary.ref (MPmap.empty : substitutive_objects MPmap.t) + ~name:"MODULE-SUBSTOBJS" + let missing_handler = ref (fun mp -> assert false) + let set_missing_handler f = (missing_handler := f) + let set mp objs = (table := MPmap.add mp objs !table) + let get mp = + try MPmap.find mp !table with Not_found -> !missing_handler mp + end -*) -type substitutive_objects = - mod_bound_id list * module_path * lib_objects +(** Some utilities about substitutive objects : + substitution, expansion *) + +let sobjs_no_functor (mbids,_) = List.is_empty mbids + +let subst_aobjs sub = function + | Objs o -> Objs (Lib.subst_objects sub o) + | Ref (mp, sub0) -> Ref (mp, join sub0 sub) + +let subst_sobjs sub (mbids,aobjs) = (mbids, subst_aobjs sub aobjs) + +let expand_aobjs = function + | Objs o -> o + | Ref (mp, sub) -> + match ModSubstObjs.get mp with + | (_,Objs o) -> Lib.subst_objects sub o + | _ -> assert false (* Invariant : any alias points to concrete objs *) +let expand_sobjs (_,aobjs) = expand_aobjs aobjs -(* For each module, we store the following things: - In modtab_substobjs: substitutive_objects - when we will do Module M:=N, the objects of N will be reloaded - with M after substitution +(** {6 ModObjs : a cache of module objects} - In modtab_objects: "substituted objects" @ "keep objects" + For each module, we also store a cache of + "prefix", "substituted objects", "keep objects". + This is used for instance to implement the "Import" command. - substituted objects - + substituted objects : roughly the objects above after the substitution - we need to keep them to call open_object when the module is opened (imported) - keep objects - + keep objects : The list of non-substitutive objects - as above, for each of them we will call open_object when the module is opened (Some) Invariants: - * If the module is a functor, the two latter lists are empty. + * If the module is a functor, it won't appear in this cache. * Module objects in substitutive_objects part have empty substituted objects. @@ -114,190 +121,88 @@ type substitutive_objects = * Modules which where created with Module M:=mexpr or with Module M:SIG. ... End M. have the keep list empty. *) -let modtab_substobjs = - ref (MPmap.empty : substitutive_objects MPmap.t) -let modtab_objects = - ref (MPmap.empty : (object_prefix * lib_objects) MPmap.t) - - -(* currently started interactive module (if any) - its arguments (if it - is a functor) and declared output type *) -let openmod_info = - ref ((MPfile(initial_dir),[],None,[]) - : module_path * mod_bound_id list * - (module_struct_entry * int option) option * - module_type_body list) - -(* The library_cache here is needed to avoid recalculations of - substituted modules object during "reloading" of libraries *) -let library_cache = ref Dirmap.empty - -let _ = Summary.declare_summary "MODULE-INFO" - { Summary.freeze_function = (fun () -> - !modtab_substobjs, - !modtab_objects, - !openmod_info, - !library_cache); - Summary.unfreeze_function = (fun (sobjs,objs,info,libcache) -> - modtab_substobjs := sobjs; - modtab_objects := objs; - openmod_info := info; - library_cache := libcache); - Summary.init_function = (fun () -> - modtab_substobjs := MPmap.empty; - modtab_objects := MPmap.empty; - openmod_info := ((MPfile(initial_dir), - [],None,[])); - library_cache := Dirmap.empty) } - -(* auxiliary functions to transform full_path and kernel_name given - by Lib into module_path and dir_path needed for modules *) -let mp_of_kn kn = - let mp,sec,l = repr_kn kn in - if sec=empty_dirpath then - MPdot (mp,l) - else - anomaly ("Non-empty section in module name!" ^ string_of_kn kn) - -let dir_of_sp sp = - let dir,id = repr_path sp in - add_dirpath_suffix dir id - -(* Subtyping checks *) - -let check_sub mtb sub_mtb_l = - (* The constraints are checked and forgot immediately : *) - ignore (List.fold_right - (fun sub_mtb env -> - Environ.add_constraints - (Subtyping.check_subtypes env mtb sub_mtb) env) - sub_mtb_l (Global.env())) +type module_objects = object_prefix * Lib.lib_objects * Lib.lib_objects -(* This function checks if the type calculated for the module [mp] is - a subtype of all signatures in [sub_mtb_l]. Uses only the global - environment. *) - -let check_subtypes mp sub_mtb_l = - let mb = - try Global.lookup_module mp - with Not_found -> assert false - in - let mtb = Modops.module_type_of_module None mb in - check_sub mtb sub_mtb_l +module ModObjs : + sig + val set : module_path -> module_objects -> unit + val get : module_path -> module_objects (* may raise Not_found *) + val all : unit -> module_objects MPmap.t + end = + struct + let table = + Summary.ref (MPmap.empty : module_objects MPmap.t) + ~name:"MODULE-OBJS" + let set mp objs = (table := MPmap.add mp objs !table) + let get mp = MPmap.find mp !table + let all () = !table + end -(* Same for module type [mp] *) -let check_subtypes_mt mp sub_mtb_l = - let mtb = - try Global.lookup_modtype mp - with Not_found -> assert false - in - check_sub mtb sub_mtb_l +(** {6 Name management} -(* Create a functor type entry *) + Auxiliary functions to transform full_path and kernel_name given + by Lib into module_path and DirPath.t needed for modules +*) -let funct_entry args m = - List.fold_right - (fun (arg_id,(arg_t,_)) mte -> MSEfunctor (arg_id,arg_t,mte)) - args m +let mp_of_kn kn = + let mp,sec,l = repr_kn kn in + assert (DirPath.is_empty sec); + MPdot (mp,l) -(* Prepare the module type list for check of subtypes *) +let dir_of_sp sp = + let dir,id = repr_path sp in + add_dirpath_suffix dir id -let build_subtypes interp_modtype mp args mtys = - List.map - (fun (m,ann) -> - let inl = inline_annot ann in - let mte = interp_modtype (Global.env()) m in - let mtb = Mod_typing.translate_module_type (Global.env()) mp inl mte in - let funct_mtb = - List.fold_right - (fun (arg_id,(arg_t,arg_inl)) mte -> - let arg_t = - Mod_typing.translate_module_type (Global.env()) - (MPbound arg_id) arg_inl arg_t - in - SEBfunctor(arg_id,arg_t,mte)) - args mtb.typ_expr - in - { mtb with typ_expr = funct_mtb }) - mtys +(** {6 Declaration of module substitutive objects} *) -(* These functions register the visibility of the module and iterates - through its components. They are called by plenty module functions *) +(** These functions register the visibility of the module and iterates + through its components. They are called by plenty of module functions *) -let compute_visibility exists what i dir dirinfo = +let consistency_checks exists dir dirinfo = if exists then - if - try Nametab.locate_dir (qualid_of_dirpath dir) = dirinfo - with Not_found -> false - then - Nametab.Exactly i - else - errorlabstrm (what^"_module") - (pr_dirpath dir ++ str " should already exist!") + let globref = + try Nametab.locate_dir (qualid_of_dirpath dir) + with Not_found -> + anomaly (pr_dirpath dir ++ str " should already exist!") + in + assert (eq_global_dir_reference globref dirinfo) else if Nametab.exists_dir dir then - errorlabstrm (what^"_module") (pr_dirpath dir ++ str " already exists") - else - Nametab.Until i -(* -let do_load_and_subst_module i dir mp substobjs keep = - let prefix = (dir,(mp,empty_dirpath)) in - let dirinfo = DirModule (dir,(mp,empty_dirpath)) in - let vis = compute_visibility false "load_and_subst" i dir dirinfo in - let objects = compute_subst_objects mp substobjs resolver in - Nametab.push_dir vis dir dirinfo; - modtab_substobjs := MPmap.add mp substobjs !modtab_substobjs; - match objects with - | Some (subst,seg) -> - let seg = load_and_subst_objects (i+1) prefix subst seg in - modtab_objects := MPmap.add mp (prefix,seg) !modtab_objects; - load_objects (i+1) prefix keep; - Some (seg@keep) - | None -> - None -*) - -let do_module exists what iter_objects i dir mp substobjs keep= - let prefix = (dir,(mp,empty_dirpath)) in - let dirinfo = DirModule (dir,(mp,empty_dirpath)) in - let vis = compute_visibility exists what i dir dirinfo in - Nametab.push_dir vis dir dirinfo; - modtab_substobjs := MPmap.add mp substobjs !modtab_substobjs; - match substobjs with - ([],mp1,objs) -> - modtab_objects := MPmap.add mp (prefix,objs@keep) !modtab_objects; - iter_objects (i+1) prefix (objs@keep) - | (mbids,_,_) -> () - -let conv_names_do_module exists what iter_objects i - (sp,kn) substobjs = - let dir,mp = dir_of_sp sp, mp_of_kn kn in - do_module exists what iter_objects i dir mp substobjs [] - -(* Interactive modules and module types cannot be recached! cache_mod* - functions can be called only once (and "end_mod*" set the flag to - false then) -*) -let cache_module ((sp,kn),substobjs) = - let dir,mp = dir_of_sp sp, mp_of_kn kn in - do_module false "cache" load_objects 1 dir mp substobjs [] - -(* When this function is called the module itself is already in the - environment. This function loads its objects only *) - -let load_module i (oname,substobjs) = - conv_names_do_module false "load" load_objects i oname substobjs - -let open_module i (oname,substobjs) = - conv_names_do_module true "open" open_objects i oname substobjs - -let subst_module (subst,(mbids,mp,objs)) = - (mbids,subst_mp subst mp, subst_objects subst objs) - -let classify_module substobjs = Substitute substobjs + anomaly (pr_dirpath dir ++ str " already exists") + +let compute_visibility exists i = + if exists then Nametab.Exactly i else Nametab.Until i + +(** Iterate some function [iter_objects] on all components of a module *) + +let do_module exists iter_objects i dir mp sobjs kobjs = + let prefix = (dir,(mp,DirPath.empty)) in + let dirinfo = DirModule prefix in + consistency_checks exists dir dirinfo; + Nametab.push_dir (compute_visibility exists i) dir dirinfo; + ModSubstObjs.set mp sobjs; + (* If we're not a functor, let's iter on the internal components *) + if sobjs_no_functor sobjs then begin + let objs = expand_sobjs sobjs in + ModObjs.set mp (prefix,objs,kobjs); + iter_objects (i+1) prefix objs; + iter_objects (i+1) prefix kobjs + end + +let do_module' exists iter_objects i ((sp,kn),sobjs) = + do_module exists iter_objects i (dir_of_sp sp) (mp_of_kn kn) sobjs [] + +(** Nota: Interactive modules and module types cannot be recached! + This used to be checked here via a flag along the substobjs. *) + +let cache_module = do_module' false Lib.load_objects 1 +let load_module = do_module' false Lib.load_objects +let open_module = do_module' true Lib.open_objects +let subst_module (subst,sobjs) = subst_sobjs subst sobjs +let classify_module sobjs = Substitute sobjs let (in_module : substitutive_objects -> obj), (out_module : obj -> substitutive_objects) = @@ -308,755 +213,748 @@ let (in_module : substitutive_objects -> obj), subst_function = subst_module; classify_function = classify_module } -let cache_keep _ = anomaly "This module should not be cached!" -let load_keep i ((sp,kn),seg) = - let mp = mp_of_kn kn in - let prefix = dir_of_sp sp, (mp,empty_dirpath) in - begin - try - let prefix',objects = MPmap.find mp !modtab_objects in - if prefix' <> prefix then - anomaly "Two different modules with the same path!"; - modtab_objects := MPmap.add mp (prefix,objects@seg) !modtab_objects; - with - Not_found -> anomaly "Keep objects before substitutive" - end; - load_objects i prefix seg - -let open_keep i ((sp,kn),seg) = - let dirpath,mp = dir_of_sp sp, mp_of_kn kn in - open_objects i (dirpath,(mp,empty_dirpath)) seg - -let in_modkeep : lib_objects -> obj = - declare_object {(default_object "MODULE KEEP OBJECTS") with +(** {6 Declaration of module keep objects} *) + +let cache_keep _ = anomaly (Pp.str "This module should not be cached!") + +let load_keep i ((sp,kn),kobjs) = + (* Invariant : seg isn't empty *) + let dir = dir_of_sp sp and mp = mp_of_kn kn in + let prefix = (dir,(mp,DirPath.empty)) in + let prefix',sobjs,kobjs0 = + try ModObjs.get mp + with Not_found -> assert false (* a substobjs should already be loaded *) + in + assert (eq_op prefix' prefix); + assert (List.is_empty kobjs0); + ModObjs.set mp (prefix,sobjs,kobjs); + Lib.load_objects i prefix kobjs + +let open_keep i ((sp,kn),kobjs) = + let dir = dir_of_sp sp and mp = mp_of_kn kn in + let prefix = (dir,(mp,DirPath.empty)) in + Lib.open_objects i prefix kobjs + +let in_modkeep : Lib.lib_objects -> obj = + declare_object {(default_object "MODULE KEEP") with cache_function = cache_keep; load_function = load_keep; open_function = open_keep } -(* we remember objects for a module type. In case of a declaration: - Module M:SIG:=... - The module M gets its objects from SIG -*) -let modtypetab = - ref (MPmap.empty : substitutive_objects MPmap.t) -(* currently started interactive module type. We remember its arguments - if it is a functor type *) -let openmodtype_info = - ref ([],[] : mod_bound_id list * module_type_body list) +(** {6 Declaration of module type substitutive objects} *) -let _ = Summary.declare_summary "MODTYPE-INFO" - { Summary.freeze_function = (fun () -> - !modtypetab,!openmodtype_info); - Summary.unfreeze_function = (fun ft -> - modtypetab := fst ft; - openmodtype_info := snd ft); - Summary.init_function = (fun () -> - modtypetab := MPmap.empty; - openmodtype_info := [],[]) } +(** Nota: Interactive modules and module types cannot be recached! + This used to be checked more properly here. *) +let do_modtype i sp mp sobjs = + if Nametab.exists_modtype sp then + anomaly (pr_path sp ++ str " already exists"); + Nametab.push_modtype (Nametab.Until i) sp mp; + ModSubstObjs.set mp sobjs -let cache_modtype ((sp,kn),(entry,modtypeobjs,sub_mty_l)) = - let mp = mp_of_kn kn in +let cache_modtype ((sp,kn),sobjs) = do_modtype 1 sp (mp_of_kn kn) sobjs +let load_modtype i ((sp,kn),sobjs) = do_modtype i sp (mp_of_kn kn) sobjs +let subst_modtype (subst,sobjs) = subst_sobjs subst sobjs +let classify_modtype sobjs = Substitute sobjs - (* We enrich the global environment *) - let _ = - match entry with - | None -> - anomaly "You must not recache interactive module types!" - | Some (mte,inl) -> - if mp <> Global.add_modtype (basename sp) mte inl then - anomaly "Kernel and Library names do not match" +let open_modtype i ((sp,kn),_) = + let mp = mp_of_kn kn in + let mp' = + try Nametab.locate_modtype (qualid_of_path sp) + with Not_found -> + anomaly (pr_path sp ++ str " should already exist!"); in + assert (ModPath.equal mp mp'); + Nametab.push_modtype (Nametab.Exactly i) sp mp - (* Using declare_modtype should lead here, where we check - that any given subtyping is indeed accurate *) - check_subtypes_mt mp sub_mty_l; +let (in_modtype : substitutive_objects -> obj), + (out_modtype : obj -> substitutive_objects) = + declare_object_full {(default_object "MODULE TYPE") with + cache_function = cache_modtype; + open_function = open_modtype; + load_function = load_modtype; + subst_function = subst_modtype; + classify_function = classify_modtype } - if Nametab.exists_modtype sp then - errorlabstrm "cache_modtype" - (pr_path sp ++ str " already exists") ; - Nametab.push_modtype (Nametab.Until 1) sp mp; +(** {6 Declaration of substitutive objects for Include} *) - modtypetab := MPmap.add mp modtypeobjs !modtypetab +let do_include do_load do_open i ((sp,kn),aobjs) = + let dir = Libnames.dirpath sp in + let mp = KerName.modpath kn in + let prefix = (dir,(mp,DirPath.empty)) in + let o = expand_aobjs aobjs in + if do_load then Lib.load_objects i prefix o; + if do_open then Lib.open_objects i prefix o +let cache_include = do_include true true 1 +let load_include = do_include true false +let open_include = do_include false true +let subst_include (subst,aobjs) = subst_aobjs subst aobjs +let classify_include aobjs = Substitute aobjs -let load_modtype i ((sp,kn),(entry,modtypeobjs,_)) = - assert (entry = None); +let (in_include : algebraic_objects -> obj), + (out_include : obj -> algebraic_objects) = + declare_object_full {(default_object "INCLUDE") with + cache_function = cache_include; + load_function = load_include; + open_function = open_include; + subst_function = subst_include; + classify_function = classify_include } - if Nametab.exists_modtype sp then - errorlabstrm "cache_modtype" - (pr_path sp ++ str " already exists") ; - Nametab.push_modtype (Nametab.Until i) sp (mp_of_kn kn); +(** {6 Handler for missing entries in ModSubstObjs} *) - modtypetab := MPmap.add (mp_of_kn kn) modtypeobjs !modtypetab +(** Since the inner of Module Types are not added by default to + the ModSubstObjs table, we compensate this by explicit traversal + of Module Types inner objects when needed. Quite a hack... *) +let mp_id mp id = MPdot (mp, Label.of_id id) -let open_modtype i ((sp,kn),(entry,_,_)) = - assert (entry = None); +let rec register_mod_objs mp (id,obj) = match object_tag obj with + | "MODULE" -> ModSubstObjs.set (mp_id mp id) (out_module obj) + | "MODULE TYPE" -> ModSubstObjs.set (mp_id mp id) (out_modtype obj) + | "INCLUDE" -> + List.iter (register_mod_objs mp) (expand_aobjs (out_include obj)) + | _ -> () - if - try Nametab.locate_modtype (qualid_of_path sp) <> (mp_of_kn kn) - with Not_found -> true - then - errorlabstrm ("open_modtype") - (pr_path sp ++ str " should already exist!"); +let handle_missing_substobjs mp = match mp with + | MPdot (mp',l) -> + let objs = expand_sobjs (ModSubstObjs.get mp') in + List.iter (register_mod_objs mp') objs; + ModSubstObjs.get mp + | _ -> + assert false (* Only inner parts of module types should be missing *) - Nametab.push_modtype (Nametab.Exactly i) sp (mp_of_kn kn) +let () = ModSubstObjs.set_missing_handler handle_missing_substobjs -let subst_modtype (subst,(entry,(mbids,mp,objs),_)) = - assert (entry = None); - (entry,(mbids,subst_mp subst mp,subst_objects subst objs),[]) -let classify_modtype (_,substobjs,_) = - Substitute (None,substobjs,[]) -type modtype_obj = - (module_struct_entry * Entries.inline) option (* will be None in vo *) - * substitutive_objects - * module_type_body list +(** {6 From module expression to substitutive objects} *) -let in_modtype : modtype_obj -> obj = - declare_object {(default_object "MODULE TYPE") with - cache_function = cache_modtype; - open_function = open_modtype; - load_function = load_modtype; - subst_function = subst_modtype; - classify_function = classify_modtype } +(** Turn a chain of [MSEapply] into the head module_path and the + list of module_path parameters (deepest param coming first). + The left part of a [MSEapply] must be either [MSEident] or + another [MSEapply]. *) -let rec replace_module_object idl (mbids,mp,lib_stack) (mbids2,mp2,objs) mp1 = - if mbids<>[] then anomaly "Unexpected functor objects"; - let rec replace_idl = function - | _,[] -> [] - | id::idl,(id',obj)::tail when id = id' -> - if object_tag obj <> "MODULE" then anomaly "MODULE expected!"; - let substobjs = - if idl = [] then - let mp' = MPdot(mp, label_of_id id) in - mbids, mp', subst_objects (map_mp mp1 mp' empty_delta_resolver) objs - else - replace_module_object idl (out_module obj) (mbids2,mp2,objs) mp - in - (id, in_module substobjs)::tail - | idl,lobj::tail -> lobj::replace_idl (idl,tail) - in - (mbids, mp, replace_idl (idl,lib_stack)) - -let discr_resolver mb = match mb.mod_type with - | SEBstruct _ -> Some mb.mod_delta - | _ -> None (* when mp is a functor *) - -(* Small function to avoid module typing during substobjs retrivial *) -let rec get_objs_modtype_application env = function -| MSEident mp -> - MPmap.find mp !modtypetab,Environ.lookup_modtype mp env,[] -| MSEapply (fexpr, MSEident mp) -> - let objs,mtb,mp_l= get_objs_modtype_application env fexpr in - objs,mtb,mp::mp_l -| MSEapply (_,mexpr) -> - Modops.error_application_to_not_path mexpr -| _ -> error "Application of a non-functor." +let get_applications mexpr = + let rec get params = function + | MEident mp -> mp, params + | MEapply (fexpr, mp) -> get (mp::params) fexpr + | MEwith _ -> error "Non-atomic functor application." + in get [] mexpr + +(** Create the substitution corresponding to some functor applications *) let rec compute_subst env mbids sign mp_l inl = match mbids,mp_l with | _,[] -> mbids,empty_subst | [],r -> error "Application of a functor with too few arguments." | mbid::mbids,mp::mp_l -> - let farg_id, farg_b, fbody_b = Modops.destr_functor env sign in + let farg_id, farg_b, fbody_b = Modops.destr_functor sign in let mb = Environ.lookup_module mp env in let mbid_left,subst = compute_subst env mbids fbody_b mp_l inl in - let resolver = match discr_resolver mb with - | None -> empty_delta_resolver - | Some mp_delta -> - Modops.inline_delta_resolver env inl mp farg_id farg_b mp_delta + let resolver = + if Modops.is_functor mb.mod_type then empty_delta_resolver + else + Modops.inline_delta_resolver env inl mp farg_id farg_b mb.mod_delta in mbid_left,join (map_mbid mbid mp resolver) subst -let rec get_modtype_substobjs env mp_from inline = function - MSEident ln -> - MPmap.find ln !modtypetab - | MSEfunctor (mbid,_,mte) -> - let (mbids, mp, objs) = get_modtype_substobjs env mp_from inline mte in - (mbid::mbids, mp, objs) - | MSEwith (mty, With_Definition _) -> - get_modtype_substobjs env mp_from inline mty - | MSEwith (mty, With_Module (idl,mp1)) -> - let substobjs = get_modtype_substobjs env mp_from inline mty in - let modobjs = MPmap.find mp1 !modtab_substobjs in - replace_module_object idl substobjs modobjs mp1 - | MSEapply (fexpr, MSEident mp) as me -> - let (mbids, mp1, objs),mtb_mp1,mp_l = - get_objs_modtype_application env me in - let mbids_left,subst = - compute_subst env mbids mtb_mp1.typ_expr (List.rev mp_l) inline - in - (mbids_left, mp1,subst_objects subst objs) - | MSEapply (_,mexpr) -> - Modops.error_application_to_not_path mexpr - -(* push names of bound modules (and their components) to Nametab *) -(* add objects associated to them *) -let process_module_bindings argids args = - let process_arg id (mbid,(mty,ann)) = - let dir = make_dirpath [id] in - let mp = MPbound mbid in - let (mbids,mp_from,objs) = - get_modtype_substobjs (Global.env()) mp (inline_annot ann) mty in - let substobjs = (mbids,mp,subst_objects - (map_mp mp_from mp empty_delta_resolver) objs)in - do_module false "start" load_objects 1 dir mp substobjs [] +(** Create the objects of a "with Module" structure. *) + +let rec replace_module_object idl mp0 objs0 mp1 objs1 = + match idl, objs0 with + | _,[] -> [] + | id::idl,(id',obj)::tail when Id.equal id id' -> + assert (object_has_tag obj "MODULE"); + let mp_id = MPdot(mp0, Label.of_id id) in + let objs = match idl with + | [] -> Lib.subst_objects (map_mp mp1 mp_id empty_delta_resolver) objs1 + | _ -> + let objs_id = expand_sobjs (out_module obj) in + replace_module_object idl mp_id objs_id mp1 objs1 in - List.iter2 process_arg argids args - -(* Same with module_type_body *) - -let rec seb2mse = function - | SEBident mp -> MSEident mp - | SEBapply (s,s',_) -> MSEapply(seb2mse s, seb2mse s') - | SEBwith (s,With_module_body (l,mp)) -> MSEwith(seb2mse s,With_Module(l,mp)) - | SEBwith (s,With_definition_body(l,cb)) -> - (match cb.const_body with - | Def c -> MSEwith(seb2mse s,With_Definition(l,Declarations.force c)) - | _ -> assert false) - | _ -> failwith "seb2mse: received a non-atomic seb" - -let process_module_seb_binding mbid seb = - process_module_bindings [id_of_mbid mbid] - [mbid, - (seb2mse seb, - { ann_inline = DefaultInline; ann_scope_subst = [] })] - -let intern_args interp_modtype (idl,(arg,ann)) = - let inl = inline_annot ann in + (id, in_module ([], Objs objs))::tail + | idl,lobj::tail -> lobj::replace_module_object idl mp0 tail mp1 objs1 + +let type_of_mod mp env = function + |true -> (Environ.lookup_module mp env).mod_type + |false -> (Environ.lookup_modtype mp env).mod_type + +let rec get_module_path = function + |MEident mp -> mp + |MEwith (me,_) -> get_module_path me + |MEapply (me,_) -> get_module_path me + +(** Substitutive objects of a module expression (or module type) *) + +let rec get_module_sobjs is_mod env inl = function + |MEident mp -> + begin match ModSubstObjs.get mp with + |(mbids,Objs _) when not (ModPath.is_bound mp) -> + (mbids,Ref (mp, empty_subst)) (* we create an alias *) + |sobjs -> sobjs + end + |MEwith (mty, WithDef _) -> get_module_sobjs is_mod env inl mty + |MEwith (mty, WithMod (idl,mp1)) -> + assert (not is_mod); + let sobjs0 = get_module_sobjs is_mod env inl mty in + assert (sobjs_no_functor sobjs0); + (* For now, we expanse everything, to be safe *) + let mp0 = get_module_path mty in + let objs0 = expand_sobjs sobjs0 in + let objs1 = expand_sobjs (ModSubstObjs.get mp1) in + ([], Objs (replace_module_object idl mp0 objs0 mp1 objs1)) + |MEapply _ as me -> + let mp1, mp_l = get_applications me in + let mbids, aobjs = get_module_sobjs is_mod env inl (MEident mp1) in + let typ = type_of_mod mp1 env is_mod in + let mbids_left,subst = compute_subst env mbids typ mp_l inl in + (mbids_left, subst_aobjs subst aobjs) + +let get_functor_sobjs is_mod env inl (params,mexpr) = + let (mbids, aobjs) = get_module_sobjs is_mod env inl mexpr in + (List.map fst params @ mbids, aobjs) + + +(** {6 Handling of module parameters} *) + +(** For printing modules, [process_module_binding] adds names of + bound module (and its components) to Nametab. It also loads + objects associated to it. *) + +let process_module_binding mbid me = + let dir = DirPath.make [MBId.to_id mbid] in + let mp = MPbound mbid in + let sobjs = get_module_sobjs false (Global.env()) (default_inline ()) me in + let subst = map_mp (get_module_path me) mp empty_delta_resolver in + let sobjs = subst_sobjs subst sobjs in + do_module false Lib.load_objects 1 dir mp sobjs [] + +(** Process a declaration of functor parameter(s) (Id1 .. Idn : Typ) + i.e. possibly multiple names with the same module type. + Global environment is updated on the fly. + Objects in these parameters are also loaded. + Output is accumulated on top of [acc] (in reverse order). *) + +let intern_arg interp_modast acc (idl,(typ,ann)) = + let inl = inl2intopt ann in let lib_dir = Lib.library_dp() in - let mbids = List.map (fun (_,id) -> make_mbid lib_dir id) idl in - let mty = interp_modtype (Global.env()) arg in - let dirs = List.map (fun (_,id) -> make_dirpath [id]) idl in - let (mbi,mp_from,objs) = get_modtype_substobjs (Global.env()) - (MPbound (List.hd mbids)) inl mty in - List.map2 - (fun dir mbid -> - let resolver = Global.add_module_parameter mbid mty inl in - let mp = MPbound mbid in - let substobjs = (mbi,mp,subst_objects - (map_mp mp_from mp resolver) objs) in - do_module false "interp" load_objects 1 dir mp substobjs []; - (mbid,(mty,inl))) - dirs mbids - -let start_module_ interp_modtype export id args res fs = - let mp = Global.start_module id in - let arg_entries = List.concat (List.map (intern_args interp_modtype) args) in - let res_entry_o, sub_body_l = match res with - | Enforce (res,ann) -> - let inl = inline_annot ann in - let mte = interp_modtype (Global.env()) res in - let _ = Mod_typing.translate_struct_type_entry (Global.env()) inl mte in - Some (mte,inl), [] - | Check resl -> - None, build_subtypes interp_modtype mp arg_entries resl - in - let mbids = List.map fst arg_entries in - openmod_info:=(mp,mbids,res_entry_o,sub_body_l); - let prefix = Lib.start_module export id mp fs in - Nametab.push_dir (Nametab.Until 1) (fst prefix) (DirOpenModule prefix); - Lib.add_frozen_state (); mp + let env = Global.env() in + let mty,_ = interp_modast env ModType typ in + let sobjs = get_module_sobjs false env inl mty in + let mp0 = get_module_path mty in + List.fold_left + (fun acc (_,id) -> + let dir = DirPath.make [id] in + let mbid = MBId.make lib_dir id in + let mp = MPbound mbid in + let resolver = Global.add_module_parameter mbid mty inl in + let sobjs = subst_sobjs (map_mp mp0 mp resolver) sobjs in + do_module false Lib.load_objects 1 dir mp sobjs []; + (mbid,mty,inl)::acc) + acc idl + +(** Process a list of declarations of functor parameters + (Id11 .. Id1n : Typ1)..(Idk1 .. Idkm : Typk) + Global environment is updated on the fly. + The calls to [interp_modast] should be interleaved with these + env updates, otherwise some "with Definition" could be rejected. + Returns a list of mbids and entries (in reversed order). + + This used to be a [List.concat (List.map ...)], but this should + be more efficient and independent of [List.map] eval order. +*) +let intern_args interp_modast params = + List.fold_left (intern_arg interp_modast) [] params -let end_module () = - let oldoname,oldprefix,fs,lib_stack = Lib.end_module () in - let mp,mbids, res_o, sub_l = !openmod_info in - let substitute, keep, special = Lib.classify_segment lib_stack in - - let mp_from,substobjs, keep, special = try - match res_o with - | None -> - (* the module is not sealed *) - None,( mbids, mp, substitute), keep, special - | Some (MSEident ln as mty, inline) -> - let (mbids1,mp1,objs) = - get_modtype_substobjs (Global.env()) mp inline mty in - Some mp1,(mbids@mbids1,mp1,objs), [], [] - | Some (MSEwith _ as mty, inline) -> - let (mbids1,mp1,objs) = - get_modtype_substobjs (Global.env()) mp inline mty in - Some mp1,(mbids@mbids1,mp1,objs), [], [] - | Some (MSEfunctor _, _) -> - anomaly "Funsig cannot be here..." - | Some (MSEapply _ as mty, inline) -> - let (mbids1,mp1,objs) = - get_modtype_substobjs (Global.env()) mp inline mty in - Some mp1,(mbids@mbids1,mp1,objs), [], [] - with - Not_found -> anomaly "Module objects not found..." - in - (* must be called after get_modtype_substobjs, because of possible - dependencies on functor arguments *) +(** {6 Auxiliary functions concerning subtyping checks} *) - let id = basename (fst oldoname) in - let mp,resolver = Global.end_module fs id res_o in +let check_sub mtb sub_mtb_l = + (* The constraints are checked and forgot immediately : *) + ignore (List.fold_right + (fun sub_mtb env -> + Environ.add_constraints + (Subtyping.check_subtypes env mtb sub_mtb) env) + sub_mtb_l (Global.env())) - check_subtypes mp sub_l; +(** This function checks if the type calculated for the module [mp] is + a subtype of all signatures in [sub_mtb_l]. Uses only the global + environment. *) -(* we substitute objects if the module is - sealed by a signature (ie. mp_from != None *) - let substobjs = match mp_from,substobjs with - None,_ -> substobjs - | Some mp_from,(mbids,_,objs) -> - (mbids,mp,subst_objects (map_mp mp_from mp resolver) objs) +let check_subtypes mp sub_mtb_l = + let mb = + try Global.lookup_module mp with Not_found -> assert false in - let node = in_module substobjs in - let objects = - if keep = [] || mbids <> [] then - special@[node] (* no keep objects or we are defining a functor *) - else - special@[node;in_modkeep keep] (* otherwise *) + let mtb = Modops.module_type_of_module mb in + check_sub mtb sub_mtb_l + +(** Same for module type [mp] *) + +let check_subtypes_mt mp sub_mtb_l = + let mtb = + try Global.lookup_modtype mp with Not_found -> assert false in - let newoname = Lib.add_leaves id objects in + check_sub mtb sub_mtb_l - if (fst newoname) <> (fst oldoname) then - anomaly "Names generated on start_ and end_module do not match"; - if mp_of_kn (snd newoname) <> mp then - anomaly "Kernel and Library names do not match"; +(** Create a params entry. + In [args], the youngest module param now comes first. *) - Lib.add_frozen_state () (* to prevent recaching *); - mp +let mk_params_entry args = + List.rev_map (fun (mbid,arg_t,_) -> (mbid,arg_t)) args +(** Create a functor type struct. + In [args], the youngest module param now comes first. *) +let mk_funct_type env args seb0 = + List.fold_left + (fun seb (arg_id,arg_t,arg_inl) -> + let mp = MPbound arg_id in + let arg_t = Mod_typing.translate_modtype env mp arg_inl ([],arg_t) in + MoreFunctor(arg_id,arg_t,seb)) + seb0 args -let module_objects mp = - let prefix,objects = MPmap.find mp !modtab_objects in - segment_of_objects prefix objects +(** Prepare the module type list for check of subtypes *) +let build_subtypes interp_modast env mp args mtys = + List.map + (fun (m,ann) -> + let inl = inl2intopt ann in + let mte,_ = interp_modast env ModType m in + let mtb = Mod_typing.translate_modtype env mp inl ([],mte) in + { mtb with mod_type = mk_funct_type env args mtb.mod_type }) + mtys -(************************************************************************) -(* libraries *) +(** {6 Current module information} -type library_name = dir_path + This information is stored by each [start_module] for use + in a later [end_module]. *) -(* The first two will form substitutive_objects, the last one is keep *) -type library_objects = - module_path * lib_objects * lib_objects +type current_module_info = { + cur_typ : (module_struct_entry * int option) option; (** type via ":" *) + cur_typs : module_type_body list (** types via "<:" *) +} +let default_module_info = { cur_typ = None; cur_typs = [] } -let register_library dir cenv objs digest = - let mp = MPfile dir in - let substobjs, keep = - try - ignore(Global.lookup_module mp); - (* if it's in the environment, the cached objects should be correct *) - Dirmap.find dir !library_cache - with Not_found -> - if mp <> Global.import cenv digest then - anomaly "Unexpected disk module name"; - let mp,substitute,keep = objs in - let substobjs = [], mp, substitute in - let modobjs = substobjs, keep in - library_cache := Dirmap.add dir modobjs !library_cache; - modobjs - in - do_module false "register_library" load_objects 1 dir mp substobjs keep +let openmod_info = Summary.ref default_module_info ~name:"MODULE-INFO" -let start_library dir = - let mp = Global.start_library dir in - openmod_info:=mp,[],None,[]; - Lib.start_compilation dir mp; - Lib.add_frozen_state () -let end_library_hook = ref ignore -let set_end_library_hook f = end_library_hook := f +(** {6 Current module type information} -let end_library dir = - !end_library_hook(); - let prefix, lib_stack = Lib.end_compilation dir in - let mp,cenv = Global.export dir in - let substitute, keep, _ = Lib.classify_segment lib_stack in - cenv,(mp,substitute,keep) + This information is stored by each [start_modtype] for use + in a later [end_modtype]. *) +let openmodtype_info = + Summary.ref ([] : module_type_body list) ~name:"MODTYPE-INFO" -(* implementation of Export M and Import M *) +(** {6 Modules : start, end, declare} *) -let really_import_module mp = - let prefix,objects = MPmap.find mp !modtab_objects in - open_objects 1 prefix objects +module RawModOps = struct +let start_module interp_modast export id args res fs = + let mp = Global.start_module id in + let arg_entries_r = intern_args interp_modast args in + let env = Global.env () in + let res_entry_o, subtyps = match res with + | Enforce (res,ann) -> + let inl = inl2intopt ann in + let mte,_ = interp_modast env ModType res in + (* We check immediately that mte is well-formed *) + let _ = Mod_typing.translate_mse env None inl mte in + Some (mte,inl), [] + | Check resl -> + None, build_subtypes interp_modast env mp arg_entries_r resl + in + openmod_info := { cur_typ = res_entry_o; cur_typs = subtyps }; + let prefix = Lib.start_module export id mp fs in + Nametab.push_dir (Nametab.Until 1) (fst prefix) (DirOpenModule prefix); + Lib.add_frozen_state (); mp -let cache_import (_,(_,mp)) = -(* for non-substitutive exports: - let mp = Nametab.locate_module (qualid_of_dirpath dir) in *) - really_import_module mp +let end_module () = + let oldoname,oldprefix,fs,lib_stack = Lib.end_module () in + let substitute, keep, special = Lib.classify_segment lib_stack in + let m_info = !openmod_info in -let classify_import (export,_ as obj) = - if export then Substitute obj else Dispose + (* For sealed modules, we use the substitutive objects of their signatures *) + let sobjs0, keep, special = match m_info.cur_typ with + | None -> ([], Objs substitute), keep, special + | Some (mty, inline) -> + get_module_sobjs false (Global.env()) inline mty, [], [] + in + let id = basename (fst oldoname) in + let mp,mbids,resolver = Global.end_module fs id m_info.cur_typ in + let sobjs = let (ms,objs) = sobjs0 in (mbids@ms,objs) in -let subst_import (subst,(export,mp as obj)) = - let mp' = subst_mp subst mp in - if mp'==mp then obj else - (export,mp') + check_subtypes mp m_info.cur_typs; -let in_import = - declare_object {(default_object "IMPORT MODULE") with - cache_function = cache_import; - open_function = (fun i o -> if i=1 then cache_import o); - subst_function = subst_import; - classify_function = classify_import } + (* We substitute objects if the module is sealed by a signature *) + let sobjs = + match m_info.cur_typ with + | None -> sobjs + | Some (mty, _) -> + subst_sobjs (map_mp (get_module_path mty) mp resolver) sobjs + in + let node = in_module sobjs in + (* We add the keep objects, if any, and if this isn't a functor *) + let objects = match keep, mbids with + | [], _ | _, _ :: _ -> special@[node] + | _ -> special@[node;in_modkeep keep] + in + let newoname = Lib.add_leaves id objects in + (* Name consistency check : start_ vs. end_module, kernel vs. library *) + assert (eq_full_path (fst newoname) (fst oldoname)); + assert (ModPath.equal (mp_of_kn (snd newoname)) mp); -let import_module export mp = - Lib.add_anonymous_leaf (in_import (export,mp)) + Lib.add_frozen_state () (* to prevent recaching *); + mp -(************************************************************************) -(* module types *) +let declare_module interp_modast id args res mexpr_o fs = + (* We simulate the beginning of an interactive module, + then we adds the module parameters to the global env. *) + let mp = Global.start_module id in + let arg_entries_r = intern_args interp_modast args in + let params = mk_params_entry arg_entries_r in + let env = Global.env () in + let mty_entry_o, subs, inl_res = match res with + | Enforce (mty,ann) -> + Some (fst (interp_modast env ModType mty)), [], inl2intopt ann + | Check mtys -> + None, build_subtypes interp_modast env mp arg_entries_r mtys, + default_inline () + in + let mexpr_entry_o, inl_expr = match mexpr_o with + | None -> None, default_inline () + | Some (mexpr,ann) -> + Some (fst (interp_modast env Module mexpr)), inl2intopt ann + in + let entry = match mexpr_entry_o, mty_entry_o with + | None, None -> assert false (* No body, no type ... *) + | None, Some typ -> MType (params, typ) + | Some body, otyp -> MExpr (params, body, otyp) + in + let sobjs, mp0 = match entry with + | MType (_,mte) | MExpr (_,_,Some mte) -> + get_functor_sobjs false env inl_res (params,mte), get_module_path mte + | MExpr (_,me,None) -> + get_functor_sobjs true env inl_expr (params,me), get_module_path me + in + (* Undo the simulated interactive building of the module + and declare the module as a whole *) + Summary.unfreeze_summaries fs; + let inl = match inl_expr with + | None -> None + | _ -> inl_res + in + let mp_env,resolver = Global.add_module id entry inl in + + (* Name consistency check : kernel vs. library *) + assert (ModPath.equal mp (mp_of_kn (Lib.make_kn id))); + assert (ModPath.equal mp mp_env); + + check_subtypes mp subs; -let start_modtype_ interp_modtype id args mtys fs = + let sobjs = subst_sobjs (map_mp mp0 mp resolver) sobjs in + ignore (Lib.add_leaf id (in_module sobjs)); + mp + +end + +(** {6 Module types : start, end, declare} *) + +module RawModTypeOps = struct + +let start_modtype interp_modast id args mtys fs = let mp = Global.start_modtype id in - let arg_entries = List.concat (List.map (intern_args interp_modtype) args) in - let sub_mty_l = build_subtypes interp_modtype mp arg_entries mtys in - let mbids = List.map fst arg_entries in - openmodtype_info := mbids, sub_mty_l; + let arg_entries_r = intern_args interp_modast args in + let env = Global.env () in + let sub_mty_l = build_subtypes interp_modast env mp arg_entries_r mtys in + openmodtype_info := sub_mty_l; let prefix = Lib.start_modtype id mp fs in Nametab.push_dir (Nametab.Until 1) (fst prefix) (DirOpenModtype prefix); Lib.add_frozen_state (); mp - let end_modtype () = let oldoname,prefix,fs,lib_stack = Lib.end_modtype () in let id = basename (fst oldoname) in let substitute, _, special = Lib.classify_segment lib_stack in - let mbids, sub_mty_l = !openmodtype_info in - let mp = Global.end_modtype fs id in - let modtypeobjs = mbids, mp, substitute in + let sub_mty_l = !openmodtype_info in + let mp, mbids = Global.end_modtype fs id in + let modtypeobjs = (mbids, Objs substitute) in check_subtypes_mt mp sub_mty_l; - let oname = Lib.add_leaves id (special@[in_modtype (None, modtypeobjs,[])]) + let oname = Lib.add_leaves id (special@[in_modtype modtypeobjs]) in - if fst oname <> fst oldoname then - anomaly - "Section paths generated on start_ and end_modtype do not match"; - if (mp_of_kn (snd oname)) <> mp then - anomaly - "Kernel and Library names do not match"; + (* Check name consistence : start_ vs. end_modtype, kernel vs. library *) + assert (eq_full_path (fst oname) (fst oldoname)); + assert (ModPath.equal (mp_of_kn (snd oname)) mp); Lib.add_frozen_state ()(* to prevent recaching *); mp - -let declare_modtype_ interp_modtype id args mtys (mty,ann) fs = - let inl = inline_annot ann in - let mmp = Global.start_modtype id in - let arg_entries = List.concat (List.map (intern_args interp_modtype) args) in - let entry = funct_entry arg_entries (interp_modtype (Global.env()) mty) in - (* NB: check of subtyping will be done in cache_modtype *) - let sub_mty_l = build_subtypes interp_modtype mmp arg_entries mtys in - let (mbids,mp_from,objs) = get_modtype_substobjs (Global.env()) mmp inl entry in - (* Undo the simulated interactive building of the module type *) - (* and declare the module type as a whole *) - - register_scope_subst ann.ann_scope_subst; - let substobjs = (mbids,mmp, - subst_objects (map_mp mp_from mmp empty_delta_resolver) objs) - in - reset_scope_subst (); +let declare_modtype interp_modast id args mtys (mty,ann) fs = + let inl = inl2intopt ann in + (* We simulate the beginning of an interactive module, + then we adds the module parameters to the global env. *) + let mp = Global.start_modtype id in + let arg_entries_r = intern_args interp_modast args in + let params = mk_params_entry arg_entries_r in + let env = Global.env () in + let entry = params, fst (interp_modast env ModType mty) in + let sub_mty_l = build_subtypes interp_modast env mp arg_entries_r mtys in + let sobjs = get_functor_sobjs false env inl entry in + let subst = map_mp (get_module_path (snd entry)) mp empty_delta_resolver in + let sobjs = subst_sobjs subst sobjs in + + (* Undo the simulated interactive building of the module type + and declare the module type as a whole *) Summary.unfreeze_summaries fs; - ignore (add_leaf id (in_modtype (Some (entry,inl), substobjs, sub_mty_l))); - mmp - - -(* Small function to avoid module typing during substobjs retrivial *) -let rec get_objs_module_application env = function -| MSEident mp -> - MPmap.find mp !modtab_substobjs,Environ.lookup_module mp env,[] -| MSEapply (fexpr, MSEident mp) -> - let objs,mtb,mp_l= get_objs_module_application env fexpr in - objs,mtb,mp::mp_l -| MSEapply (_,mexpr) -> - Modops.error_application_to_not_path mexpr -| _ -> error "Application of a non-functor." - - -let rec get_module_substobjs env mp_from inl = function - | MSEident mp -> MPmap.find mp !modtab_substobjs - | MSEfunctor (mbid,mty,mexpr) -> - let (mbids, mp, objs) = get_module_substobjs env mp_from inl mexpr in - (mbid::mbids, mp, objs) - | MSEapply (fexpr, MSEident mp) as me -> - let (mbids, mp1, objs),mb_mp1,mp_l = - get_objs_module_application env me - in - let mbids_left,subst = - compute_subst env mbids mb_mp1.mod_type (List.rev mp_l) inl in - (mbids_left, mp1,subst_objects subst objs) - | MSEapply (_,mexpr) -> Modops.error_application_to_not_path mexpr - | MSEwith (mty, With_Definition _) -> get_module_substobjs env mp_from inl mty - | MSEwith (mty, With_Module (idl,mp)) -> assert false - - -let declare_module_ interp_modtype interp_modexpr id args res mexpr_o fs = - let mmp = Global.start_module id in - let arg_entries = List.concat (List.map (intern_args interp_modtype) args) in - - let funct f m = funct_entry arg_entries (f (Global.env ()) m) in - let env = Global.env() in - let mty_entry_o, subs, inl_res = match res with - | Enforce (mty,ann) -> - Some (funct interp_modtype mty), [], inline_annot ann - | Check mtys -> - None, build_subtypes interp_modtype mmp arg_entries mtys, - default_inline () - in - - (*let subs = List.map (Mod_typing.translate_module_type env mmp) mty_sub_l in *) - let mexpr_entry_o, inl_expr, scl = match mexpr_o with - | None -> None, default_inline (), [] - | Some (mexpr,ann) -> - Some (funct interp_modexpr mexpr), inline_annot ann, ann.ann_scope_subst - in - let entry = - {mod_entry_type = mty_entry_o; - mod_entry_expr = mexpr_entry_o } - in + (* We enrich the global environment *) + let mp_env = Global.add_modtype id entry inl in - let substobjs = - match entry with - | {mod_entry_type = Some mte} -> get_modtype_substobjs env mmp inl_res mte - | {mod_entry_expr = Some mexpr} -> get_module_substobjs env mmp inl_expr mexpr - | _ -> anomaly "declare_module: No type, no body ..." - in - let (mbids,mp_from,objs) = substobjs in - (* Undo the simulated interactive building of the module *) - (* and declare the module as a whole *) - Summary.unfreeze_summaries fs; - let mp = mp_of_kn (Lib.make_kn id) in - let inl = if inl_expr = None then None else inl_res in (*PLTODO *) - let mp_env,resolver = Global.add_module id entry inl in + (* Name consistency check : kernel vs. library *) + assert (ModPath.equal mp_env mp); - if mp_env <> mp then anomaly "Kernel and Library names do not match"; + (* Subtyping checks *) + check_subtypes_mt mp sub_mty_l; - - check_subtypes mp subs; - register_scope_subst scl; - let substobjs = (mbids,mp_env, - subst_objects(map_mp mp_from mp_env resolver) objs) in - reset_scope_subst (); - ignore (add_leaf - id - (in_module substobjs)); - mmp - -(* Include *) - -let rec subst_inc_expr subst me = - match me with - | MSEident mp -> MSEident (subst_mp subst mp) - | MSEwith (me,With_Module(idl,mp)) -> - MSEwith (subst_inc_expr subst me, - With_Module(idl,subst_mp subst mp)) - | MSEwith (me,With_Definition(idl,const))-> - let const1 = Mod_subst.from_val const in - let force = Mod_subst.force subst_mps in - MSEwith (subst_inc_expr subst me, - With_Definition(idl,force (subst_substituted - subst const1))) - | MSEapply (me1,me2) -> - MSEapply (subst_inc_expr subst me1, - subst_inc_expr subst me2) - | MSEfunctor(mbid,me1,me2) -> - MSEfunctor (mbid, subst_inc_expr subst me1, subst_inc_expr subst me2) - -let lift_oname (sp,kn) = - let mp,_,_ = Names.repr_kn kn in - let dir,_ = Libnames.repr_path sp in - (dir,mp) - -let cache_include (oname,(me,(mbis,mp1,objs))) = - let dir,mp1 = lift_oname oname in - let prefix = (dir,(mp1,empty_dirpath)) in - load_objects 1 prefix objs; - open_objects 1 prefix objs - -let load_include i (oname,(me,(mbis,mp1,objs))) = - let dir,mp1 = lift_oname oname in - let prefix = (dir,(mp1,empty_dirpath)) in - load_objects i prefix objs - -let open_include i (oname,(me,(mbis,mp1,objs))) = - let dir,mp1 = lift_oname oname in - let prefix = (dir,(mp1,empty_dirpath)) in - open_objects i prefix objs - -let subst_include (subst,(me,substobj)) = - let (mbids,mp,objs) = substobj in - let substobjs = (mbids,subst_mp subst mp,subst_objects subst objs) in - (subst_inc_expr subst me,substobjs) - -let classify_include (me,substobjs) = Substitute (me,substobjs) - -type include_obj = module_struct_entry * substitutive_objects - -let (in_include : include_obj -> obj), - (out_include : obj -> include_obj) = - declare_object_full {(default_object "INCLUDE") with - cache_function = cache_include; - load_function = load_include; - open_function = open_include; - subst_function = subst_include; - classify_function = classify_include } + ignore (Lib.add_leaf id (in_modtype sobjs)); + mp -let rec include_subst env mb mbids sign inline = - match mbids with - | [] -> empty_subst - | mbid::mbids -> - let farg_id, farg_b, fbody_b = Modops.destr_functor env sign in - let subst = include_subst env mb mbids fbody_b inline in - let mp_delta = - Modops.inline_delta_resolver env inline mb.mod_mp - farg_id farg_b mb.mod_delta - in - join (map_mbid mbid mb.mod_mp mp_delta) subst - -exception NothingToDo - -let get_includeself_substobjs env objs me is_mod inline = - try - let mb_mp = match me with - | MSEident mp -> - if is_mod then - Environ.lookup_module mp env - else - Modops.module_body_of_type mp (Environ.lookup_modtype mp env) - | MSEapply(fexpr, MSEident p) as mexpr -> - let _,mb_mp,mp_l = - if is_mod then - get_objs_module_application env mexpr - else - let o,mtb_mp,mp_l = get_objs_modtype_application env mexpr in - o,Modops.module_body_of_type mtb_mp.typ_mp mtb_mp,mp_l - in - List.fold_left - (fun mb _ -> - match mb.mod_type with - | SEBfunctor(_,_,str) -> {mb with mod_type = str} - | _ -> error "Application of a functor with too much arguments.") - mb_mp mp_l - | _ -> raise NothingToDo +end + +(** {6 Include} *) + +module RawIncludeOps = struct + +let rec include_subst env mp reso mbids sign inline = match mbids with + | [] -> empty_subst + | mbid::mbids -> + let farg_id, farg_b, fbody_b = Modops.destr_functor sign in + let subst = include_subst env mp reso mbids fbody_b inline in + let mp_delta = + Modops.inline_delta_resolver env inline mp farg_id farg_b reso in - let (mbids,mp_self,objects) = objs in - let mb = Global.pack_module() in - let subst = include_subst env mb mbids mb_mp.mod_type inline in - ([],mp_self,subst_objects subst objects) - with NothingToDo -> objs + join (map_mbid mbid mp mp_delta) subst +let rec decompose_functor mpl typ = + match mpl, typ with + | [], _ -> typ + | _::mpl, MoreFunctor(_,_,str) -> decompose_functor mpl str + | _ -> error "Application of a functor with too much arguments." +exception NoIncludeSelf +let type_of_incl env is_mod = function + |MEident mp -> type_of_mod mp env is_mod + |MEapply _ as me -> + let mp0, mp_l = get_applications me in + decompose_functor mp_l (type_of_mod mp0 env is_mod) + |MEwith _ -> raise NoIncludeSelf -let declare_one_include_inner annot (me,is_mod) = +let declare_one_include interp_modast (me_ast,annot) = let env = Global.env() in - let mp1,_ = current_prefix () in - let inl = inline_annot annot in - let (mbids,mp,objs)= - if is_mod then - get_module_substobjs env mp1 inl me - else - get_modtype_substobjs env mp1 inl me in - let (mbids,mp,objs) = - if mbids <> [] then - get_includeself_substobjs env (mbids,mp,objs) me is_mod inl - else - (mbids,mp,objs) in - let id = current_mod_id() in - let resolver = Global.add_include me is_mod inl in - register_scope_subst annot.ann_scope_subst; - let substobjs = (mbids,mp1, - subst_objects (map_mp mp mp1 resolver) objs) in - reset_scope_subst (); - ignore (add_leaf id (in_include (me, substobjs))) - -let declare_one_include interp_struct (me_ast,annot) = - declare_one_include_inner annot - (interp_struct (Global.env()) me_ast) - -let declare_include_ interp_struct me_asts = - List.iter (declare_one_include interp_struct) me_asts - -(** Versions of earlier functions taking care of the freeze/unfreeze - of summaries *) + let me,kind = interp_modast env ModAny me_ast in + let is_mod = (kind == Module) in + let cur_mp = Lib.current_mp () in + let inl = inl2intopt annot in + let mbids,aobjs = get_module_sobjs is_mod env inl me in + let subst_self = + try + if List.is_empty mbids then raise NoIncludeSelf; + let typ = type_of_incl env is_mod me in + let reso,_ = Safe_typing.delta_of_senv (Global.safe_env ()) in + include_subst env cur_mp reso mbids typ inl + with NoIncludeSelf -> empty_subst + in + let base_mp = get_module_path me in + let resolver = Global.add_include me is_mod inl in + let subst = join subst_self (map_mp base_mp cur_mp resolver) in + let aobjs = subst_aobjs subst aobjs in + ignore (Lib.add_leaf (Lib.current_mod_id ()) (in_include aobjs)) + +let declare_include interp me_asts = + List.iter (declare_one_include interp) me_asts + +end + + +(** {6 Module operations handling summary freeze/unfreeze} *) let protect_summaries f = - let fs = Summary.freeze_summaries () in + let fs = Summary.freeze_summaries ~marshallable:`No in try f fs with reraise -> (* Something wrong: undo the whole process *) - Summary.unfreeze_summaries fs; raise reraise + let reraise = Errors.push reraise in + let () = Summary.unfreeze_summaries fs in + iraise reraise + +let start_module interp export id args res = + protect_summaries (RawModOps.start_module interp export id args res) -let declare_include interp_struct me_asts = - protect_summaries - (fun _ -> declare_include_ interp_struct me_asts) +let end_module = RawModOps.end_module -let declare_modtype interp_mt interp_mix id args mtys mty_l = +let declare_module interp id args mtys me_l = + let declare_me fs = match me_l with + | [] -> RawModOps.declare_module interp id args mtys None fs + | [me] -> RawModOps.declare_module interp id args mtys (Some me) fs + | me_l -> + ignore (RawModOps.start_module interp None id args mtys fs); + RawIncludeOps.declare_include interp me_l; + RawModOps.end_module () + in + protect_summaries declare_me + +let start_modtype interp id args mtys = + protect_summaries (RawModTypeOps.start_modtype interp id args mtys) + +let end_modtype = RawModTypeOps.end_modtype + +let declare_modtype interp id args mtys mty_l = let declare_mt fs = match mty_l with | [] -> assert false - | [mty] -> declare_modtype_ interp_mt id args mtys mty fs + | [mty] -> RawModTypeOps.declare_modtype interp id args mtys mty fs | mty_l -> - ignore (start_modtype_ interp_mt id args mtys fs); - declare_include_ interp_mix mty_l; - end_modtype () + ignore (RawModTypeOps.start_modtype interp id args mtys fs); + RawIncludeOps.declare_include interp mty_l; + RawModTypeOps.end_modtype () in protect_summaries declare_mt -let start_modtype interp_modtype id args mtys = - protect_summaries (start_modtype_ interp_modtype id args mtys) +let declare_include interp me_asts = + protect_summaries (fun _ -> RawIncludeOps.declare_include interp me_asts) -let declare_module interp_mt interp_me interp_mix id args mtys me_l = - let declare_me fs = match me_l with - | [] -> declare_module_ interp_mt interp_me id args mtys None fs - | [me] -> declare_module_ interp_mt interp_me id args mtys (Some me) fs - | me_l -> - ignore (start_module_ interp_mt None id args mtys fs); - declare_include_ interp_mix me_l; - end_module () + +(** {6 Libraries} *) + +type library_name = DirPath.t + +(** A library object is made of some substitutive objects + and some "keep" objects. *) + +type library_objects = Lib.lib_objects * Lib.lib_objects + +(** For the native compiler, we cache the library values *) + +type library_values = Nativecode.symbol array +let library_values = + Summary.ref (Dirmap.empty : library_values Dirmap.t) ~name:"LIBVALUES" + +let register_library dir cenv (objs:library_objects) digest univ = + let mp = MPfile dir in + let () = + try + (* Is this library already loaded ? *) + ignore(Global.lookup_module mp); + with Not_found -> + (* If not, let's do it now ... *) + let mp', values = Global.import cenv univ digest in + if not (ModPath.equal mp mp') then + anomaly (Pp.str "Unexpected disk module name"); + library_values := Dirmap.add dir values !library_values in - protect_summaries declare_me + let sobjs,keepobjs = objs in + do_module false Lib.load_objects 1 dir mp ([],Objs sobjs) keepobjs + +let get_library_symbols_tbl dir = Dirmap.find dir !library_values + +let start_library dir = + let mp = Global.start_library dir in + openmod_info := default_module_info; + Lib.start_compilation dir mp; + Lib.add_frozen_state () + +let end_library ?except dir = + let oname = Lib.end_compilation_checks dir in + let mp,cenv,ast = Global.export ?except dir in + let prefix, lib_stack = Lib.end_compilation oname in + assert (ModPath.equal mp (MPfile dir)); + let substitute, keep, _ = Lib.classify_segment lib_stack in + cenv,(substitute,keep),ast + + + +(** {6 Implementation of Import and Export commands} *) + +let really_import_module mp = + (* May raise Not_found for unknown module and for functors *) + let prefix,sobjs,keepobjs = ModObjs.get mp in + Lib.open_objects 1 prefix sobjs; + Lib.open_objects 1 prefix keepobjs + +let cache_import (_,(_,mp)) = really_import_module mp -let start_module interp_modtype export id args res = - protect_summaries (start_module_ interp_modtype export id args res) +let open_import i obj = + if Int.equal i 1 then cache_import obj + +let classify_import (export,_ as obj) = + if export then Substitute obj else Dispose + +let subst_import (subst,(export,mp as obj)) = + let mp' = subst_mp subst mp in + if mp'==mp then obj else (export,mp') + +let in_import : bool * module_path -> obj = + declare_object {(default_object "IMPORT MODULE") with + cache_function = cache_import; + open_function = open_import; + subst_function = subst_import; + classify_function = classify_import } + +let import_module export mp = + Lib.add_anonymous_leaf (in_import (export,mp)) -(*s Iterators. *) +(** {6 Iterators} *) let iter_all_segments f = - let _ = - MPmap.iter - (fun _ (prefix,objects) -> - let rec apply_obj (id,obj) = match object_tag obj with - | "INCLUDE" -> - let (_,(_,_,objs)) = out_include obj in - List.iter apply_obj objs - - | _ -> f (make_oname prefix id) obj in - List.iter apply_obj objects) - !modtab_objects + let rec apply_obj prefix (id,obj) = match object_tag obj with + | "INCLUDE" -> + let objs = expand_aobjs (out_include obj) in + List.iter (apply_obj prefix) objs + | _ -> f (make_oname prefix id) obj in - let rec apply_node = function - | sp, Leaf o -> f sp o + let apply_mod_obj _ (prefix,substobjs,keepobjs) = + List.iter (apply_obj prefix) substobjs; + List.iter (apply_obj prefix) keepobjs + in + let apply_node = function + | sp, Lib.Leaf o -> f sp o | _ -> () in - List.iter apply_node (Lib.contents_after None) + MPmap.iter apply_mod_obj (ModObjs.all ()); + List.iter apply_node (Lib.contents ()) + + +(** {6 Some types used to shorten declaremods.mli} *) + +type 'modast module_interpretor = + Environ.env -> Misctypes.module_kind -> 'modast -> + Entries.module_struct_entry * Misctypes.module_kind + +type 'modast module_params = + (Id.t Loc.located list * ('modast * inline)) list + +(** {6 Debug} *) let debug_print_modtab _ = let pr_seg = function | [] -> str "[]" | l -> str ("[." ^ string_of_int (List.length l) ^ ".]") in - let pr_modinfo mp (prefix,objects) s = + let pr_modinfo mp (prefix,substobjs,keepobjs) s = s ++ str (string_of_mp mp) ++ (spc ()) - ++ (pr_seg (segment_of_objects prefix objects)) + ++ (pr_seg (Lib.segment_of_objects prefix (substobjs@keepobjs))) in - let modules = MPmap.fold pr_modinfo !modtab_objects (mt ()) in - hov 0 modules + let modules = MPmap.fold pr_modinfo (ModObjs.all ()) (mt ()) in + hov 0 modules diff --git a/library/declaremods.mli b/library/declaremods.mli index 5019b659..c3578ec4 100644 --- a/library/declaremods.mli +++ b/library/declaremods.mli @@ -1,77 +1,42 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* string - -(** Which inline annotations should we honor, either None or the ones - whose level is less or equal to the given integer *) - -type inline = - | NoInline - | DefaultInline - | InlineAt of int - -(** The type of annotations for functor applications *) - -type funct_app_annot = - { ann_inline : inline; - ann_scope_subst : scope_subst } - -type 'a annotated = ('a * funct_app_annot) +open Vernacexpr (** {6 Modules } *) -(** [declare_module interp_modtype interp_modexpr id fargs typ expr] - declares module [id], with type constructed by [interp_modtype] - from functor arguments [fargs] and [typ] and with module body - constructed by [interp_modtype] from functor arguments [fargs] and - by [interp_modexpr] from [expr]. At least one of [typ], [expr] must - be non-empty. +type 'modast module_interpretor = + Environ.env -> Misctypes.module_kind -> 'modast -> + Entries.module_struct_entry * Misctypes.module_kind + +type 'modast module_params = + (Id.t Loc.located list * ('modast * inline)) list - The [bool] in [typ] tells if the module must be abstracted [true] - with respect to the module type or merely matched without any - restriction [false]. -*) +(** [declare_module interp_modast id fargs typ exprs] + declares module [id], with structure constructed by [interp_modast] + from functor arguments [fargs], with final type [typ]. + [exprs] is usually of length 1 (Module definition with a concrete + body), but it could also be empty ("Declare Module", with non-empty [typ]), + or multiple (body of the shape M <+ N <+ ...). *) val declare_module : - (env -> 'modast -> module_struct_entry) -> - (env -> 'modast -> module_struct_entry) -> - (env -> 'modast -> module_struct_entry * bool) -> - identifier -> - (identifier located list * ('modast annotated)) list -> - ('modast annotated) module_signature -> - ('modast annotated) list -> module_path - -val start_module : (env -> 'modast -> module_struct_entry) -> - bool option -> identifier -> - (identifier located list * ('modast annotated)) list -> - ('modast annotated) module_signature -> module_path + 'modast module_interpretor -> + Id.t -> + 'modast module_params -> + ('modast * inline) module_signature -> + ('modast * inline) list -> module_path + +val start_module : + 'modast module_interpretor -> + bool option -> Id.t -> + 'modast module_params -> + ('modast * inline) module_signature -> module_path val end_module : unit -> module_path @@ -79,49 +44,49 @@ val end_module : unit -> module_path (** {6 Module types } *) -val declare_modtype : (env -> 'modast -> module_struct_entry) -> - (env -> 'modast -> module_struct_entry * bool) -> - identifier -> - (identifier located list * ('modast annotated)) list -> - ('modast annotated) list -> - ('modast annotated) list -> +(** [declare_modtype interp_modast id fargs typs exprs] + Similar to [declare_module], except that the types could be multiple *) + +val declare_modtype : + 'modast module_interpretor -> + Id.t -> + 'modast module_params -> + ('modast * inline) list -> + ('modast * inline) list -> module_path -val start_modtype : (env -> 'modast -> module_struct_entry) -> - identifier -> (identifier located list * ('modast annotated)) list -> - ('modast annotated) list -> module_path +val start_modtype : + 'modast module_interpretor -> + Id.t -> + 'modast module_params -> + ('modast * inline) list -> module_path val end_modtype : unit -> module_path -(** {6 ... } *) -(** Objects of a module. They come in two lists: the substitutive ones - and the other *) - -val module_objects : module_path -> library_segment - - (** {6 Libraries i.e. modules on disk } *) -type library_name = dir_path +type library_name = DirPath.t type library_objects val register_library : library_name -> - Safe_typing.compiled_library -> library_objects -> Digest.t -> unit + Safe_typing.compiled_library -> library_objects -> Safe_typing.vodigest -> + Univ.universe_context_set -> unit + +val get_library_symbols_tbl : library_name -> Nativecode.symbol array val start_library : library_name -> unit val end_library : - library_name -> Safe_typing.compiled_library * library_objects - -(** set a function to be executed at end_library *) -val set_end_library_hook : (unit -> unit) -> unit + ?except:Future.UUIDSet.t -> library_name -> + Safe_typing.compiled_library * library_objects * Safe_typing.native_library (** [really_import_module mp] opens the module [mp] (in a Caml sense). It modifies Nametab and performs the [open_object] function for - every object of the module. *) + every object of the module. Raises [Not_found] when [mp] is unknown + or when [mp] corresponds to a functor. *) val really_import_module : module_path -> unit @@ -133,8 +98,8 @@ val import_module : bool -> module_path -> unit (** Include *) -val declare_include : (env -> 'struct_expr -> module_struct_entry * bool) -> - ('struct_expr annotated) list -> unit +val declare_include : + 'modast module_interpretor -> ('modast * inline) list -> unit (** {6 ... } *) (** [iter_all_segments] iterate over all segments, the modules' @@ -142,17 +107,16 @@ val declare_include : (env -> 'struct_expr -> module_struct_entry * bool) -> in an arbitrary order. The given function is applied to all leaves (together with their section path). *) -val iter_all_segments : (object_name -> obj -> unit) -> unit +val iter_all_segments : + (Libnames.object_name -> Libobject.obj -> unit) -> unit val debug_print_modtab : unit -> Pp.std_ppcmds -(*i val debug_print_modtypetab : unit -> Pp.std_ppcmds i*) - -(** For translator *) -val process_module_bindings : module_ident list -> - (mod_bound_id * (module_struct_entry annotated)) list -> unit +(** For printing modules, [process_module_binding] adds names of + bound module (and its components) to Nametab. It also loads + objects associated to it. It may raise a [Failure] when the + bound module hasn't an atomic type. *) -(** For Printer *) -val process_module_seb_binding : - mod_bound_id -> Declarations.struct_expr_body -> unit +val process_module_binding : + MBId.t -> Declarations.module_alg_expr -> unit diff --git a/library/decls.ml b/library/decls.ml index a20d438f..8d5085f7 100644 --- a/library/decls.ml +++ b/library/decls.ml @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* !vartab); - Summary.unfreeze_function = (fun ft -> vartab := ft); - Summary.init_function = (fun () -> vartab := Idmap.empty) } +let add_variable_data id o = vartab := Id.Map.add id o !vartab -let add_variable_data id o = vartab := Idmap.add id o !vartab - -let variable_path id = let (p,_,_,_) = Idmap.find id !vartab in p -let variable_opacity id = let (_,opaq,_,_) = Idmap.find id !vartab in opaq -let variable_kind id = let (_,_,_,k) = Idmap.find id !vartab in k -let variable_constraints id = let (_,_,cst,_) = Idmap.find id !vartab in cst +let variable_path id = let (p,_,_,_,_) = Id.Map.find id !vartab in p +let variable_opacity id = let (_,opaq,_,_,_) = Id.Map.find id !vartab in opaq +let variable_kind id = let (_,_,_,_,k) = Id.Map.find id !vartab in k +let variable_context id = let (_,_,ctx,_,_) = Id.Map.find id !vartab in ctx +let variable_polymorphic id = let (_,_,_,p,_) = Id.Map.find id !vartab in p let variable_secpath id = let dir = drop_dirpath_prefix (Lib.library_dp()) (variable_path id) in make_qualid dir id -let variable_exists id = Idmap.mem id !vartab +let variable_exists id = Id.Map.mem id !vartab (** Datas associated to global parameters and constants *) -let csttab = ref (Cmap.empty : logical_kind Cmap.t) - -let _ = Summary.declare_summary "CONSTANT" - { Summary.freeze_function = (fun () -> !csttab); - Summary.unfreeze_function = (fun ft -> csttab := ft); - Summary.init_function = (fun () -> csttab := Cmap.empty) } +let csttab = Summary.ref (Cmap.empty : logical_kind Cmap.t) ~name:"CONSTANT" let add_constant_kind kn k = csttab := Cmap.add kn k !csttab @@ -65,7 +57,7 @@ let initialize_named_context_for_proof () = let last_section_hyps dir = fold_named_context (fun (id,_,_) sec_ids -> - try if dir=variable_path id then id::sec_ids else sec_ids + try if DirPath.equal dir (variable_path id) then id::sec_ids else sec_ids with Not_found -> sec_ids) (Environ.named_context (Global.env())) ~init:[] diff --git a/library/decls.mli b/library/decls.mli index 6cd7887e..ac0d907d 100644 --- a/library/decls.mli +++ b/library/decls.mli @@ -1,13 +1,12 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* variable_data -> unit -val variable_path : variable -> dir_path +val variable_path : variable -> DirPath.t val variable_secpath : variable -> qualid val variable_kind : variable -> logical_kind val variable_opacity : variable -> bool -val variable_constraints : variable -> Univ.constraints +val variable_context : variable -> Univ.universe_context_set +val variable_polymorphic : variable -> polymorphic val variable_exists : variable -> bool (** Registration and access to the table of constants *) @@ -40,4 +40,4 @@ val initialize_named_context_for_proof : unit -> Environ.named_context_val (** Miscellaneous functions *) -val last_section_hyps : dir_path -> identifier list +val last_section_hyps : DirPath.t -> Id.t list diff --git a/library/dischargedhypsmap.ml b/library/dischargedhypsmap.ml index f95b6c03..e4280334 100644 --- a/library/dischargedhypsmap.ml +++ b/library/dischargedhypsmap.ml @@ -1,47 +1,19 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* - [] - -(*s Registration as global tables and rollback. *) - -let init () = - discharged_hyps_map := Spmap.empty - -let freeze () = !discharged_hyps_map - -let unfreeze dhm = discharged_hyps_map := dhm - -let _ = - Summary.declare_summary "discharged_hypothesis" - { Summary.freeze_function = freeze; - Summary.unfreeze_function = unfreeze; - Summary.init_function = init } + try Spmap.find sp !discharged_hyps_map with Not_found -> [] diff --git a/library/dischargedhypsmap.mli b/library/dischargedhypsmap.mli index 03405c32..73689201 100644 --- a/library/dischargedhypsmap.mli +++ b/library/dischargedhypsmap.mli @@ -1,20 +1,17 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* discharged_hyps -> unit val get_discharged_hyps : full_path -> discharged_hyps diff --git a/library/global.ml b/library/global.ml index 96e522a7..875097e4 100644 --- a/library/global.ml +++ b/library/global.ml @@ -1,180 +1,258 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* Safe_typing.safe_environment + val set_safe_env : Safe_typing.safe_environment -> unit + val join_safe_environment : ?except:Future.UUIDSet.t -> unit -> unit -let env_is_empty () = is_empty !global_env +end = struct -let _ = - declare_summary "Global environment" - { freeze_function = (fun () -> !global_env); +let global_env = ref Safe_typing.empty_environment + +let join_safe_environment ?except () = + global_env := Safe_typing.join_safe_environment ?except !global_env + +let () = + Summary.declare_summary global_env_summary_name + { Summary.freeze_function = (function + | `Yes -> join_safe_environment (); !global_env + | `No -> !global_env + | `Shallow -> !global_env); unfreeze_function = (fun fr -> global_env := fr); - init_function = (fun () -> global_env := empty_environment) } + init_function = (fun () -> global_env := Safe_typing.empty_environment) } -(* Then we export the functions of [Typing] on that environment. *) +let assert_not_parsing () = + if !Flags.we_are_parsing then + Errors.anomaly ( + Pp.strbrk"The global environment cannot be accessed during parsing") -let universes () = universes (env()) -let named_context () = named_context (env()) -let named_context_val () = named_context_val (env()) +let safe_env () = assert_not_parsing(); !global_env -let push_named_assum a = - let (cst,env) = push_named_assum a !global_env in - global_env := env; - cst -let push_named_def d = - let (cst,env) = push_named_def d !global_env in - global_env := env; - cst +let set_safe_env e = global_env := e +end -let add_thing add dir id thing = - let kn, newenv = add dir (label_of_id id) thing !global_env in - global_env := newenv; - kn +let safe_env = GlobalSafeEnv.safe_env +let join_safe_environment ?except () = + GlobalSafeEnv.join_safe_environment ?except () -let add_constant = add_thing add_constant -let add_mind = add_thing add_mind -let add_modtype x y inl = add_thing (fun _ x y -> add_modtype x y inl) () x y +let env () = Safe_typing.env_of_safe_env (safe_env ()) +let env_is_initial () = Safe_typing.is_initial (safe_env ()) -let add_module id me inl = - let l = label_of_id id in - let mp,resolve,new_env = add_module l me inl !global_env in - global_env := new_env; - mp,resolve - +(** Turn ops over the safe_environment state monad to ops on the global env *) -let add_constraints c = global_env := add_constraints c !global_env +let globalize0 f = GlobalSafeEnv.set_safe_env (f (safe_env ())) -let set_engagement c = global_env := set_engagement c !global_env +let globalize f = + let res,env = f (safe_env ()) in GlobalSafeEnv.set_safe_env env; res -let add_include me is_module inl = - let resolve,newenv = add_include me is_module inl !global_env in - global_env := newenv; - resolve +let globalize_with_summary fs f = + let res,env = f (safe_env ()) in + Summary.unfreeze_summaries fs; + GlobalSafeEnv.set_safe_env env; + res -let start_module id = - let l = label_of_id id in - let mp,newenv = start_module l !global_env in - global_env := newenv; - mp +(** [Safe_typing] operations, now operating on the global environment *) -let end_module fs id mtyo = - let l = label_of_id id in - let mp,resolve,newenv = end_module l mtyo !global_env in - Summary.unfreeze_summaries fs; - global_env := newenv; - mp,resolve +let i2l = Label.of_id +let push_named_assum a = globalize0 (Safe_typing.push_named_assum a) +let push_named_def d = globalize0 (Safe_typing.push_named_def d) +let add_constraints c = globalize0 (Safe_typing.add_constraints c) +let push_context_set c = globalize0 (Safe_typing.push_context_set c) +let push_context c = globalize0 (Safe_typing.push_context c) -let add_module_parameter mbid mte inl = - let resolve,newenv = add_module_parameter mbid mte inl !global_env in - global_env := newenv; - resolve +let set_engagement c = globalize0 (Safe_typing.set_engagement c) +let set_type_in_type () = globalize0 (Safe_typing.set_type_in_type) +let add_constant dir id d = globalize (Safe_typing.add_constant dir (i2l id) d) +let add_mind dir id mie = globalize (Safe_typing.add_mind dir (i2l id) mie) +let add_modtype id me inl = globalize (Safe_typing.add_modtype (i2l id) me inl) +let add_module id me inl = globalize (Safe_typing.add_module (i2l id) me inl) +let add_include me ismod inl = globalize (Safe_typing.add_include me ismod inl) +let start_module id = globalize (Safe_typing.start_module (i2l id)) +let start_modtype id = globalize (Safe_typing.start_modtype (i2l id)) -let start_modtype id = - let l = label_of_id id in - let mp,newenv = start_modtype l !global_env in - global_env := newenv; - mp +let end_module fs id mtyo = + globalize_with_summary fs (Safe_typing.end_module (i2l id) mtyo) let end_modtype fs id = - let l = label_of_id id in - let kn,newenv = end_modtype l !global_env in - Summary.unfreeze_summaries fs; - global_env := newenv; - kn + globalize_with_summary fs (Safe_typing.end_modtype (i2l id)) -let pack_module () = - pack_module !global_env +let add_module_parameter mbid mte inl = + globalize (Safe_typing.add_module_parameter mbid mte inl) +(** Queries on the global environment *) +let universes () = universes (env()) +let named_context () = named_context (env()) +let named_context_val () = named_context_val (env()) let lookup_named id = lookup_named id (env()) let lookup_constant kn = lookup_constant kn (env()) let lookup_inductive ind = Inductive.lookup_mind_specif (env()) ind +let lookup_pinductive (ind,_) = Inductive.lookup_mind_specif (env()) ind let lookup_mind kn = lookup_mind kn (env()) let lookup_module mp = lookup_module mp (env()) let lookup_modtype kn = lookup_modtype kn (env()) +let exists_objlabel id = Safe_typing.exists_objlabel id (safe_env ()) + +let opaque_tables () = Environ.opaque_tables (env ()) +let body_of_constant_body cb = Declareops.body_of_constant (opaque_tables ()) cb +let body_of_constant cst = body_of_constant_body (lookup_constant cst) +let constraints_of_constant_body cb = + Declareops.constraints_of_constant (opaque_tables ()) cb +let universes_of_constant_body cb = + Declareops.universes_of_constant (opaque_tables ()) cb + +(** Operations on kernel names *) + let constant_of_delta_kn kn = - let resolver,resolver_param = (delta_of_senv !global_env) in + let resolver,resolver_param = Safe_typing.delta_of_senv (safe_env ()) + in (* TODO : are resolver and resolver_param orthogonal ? the effect of resolver is lost if resolver_param isn't trivial at that spot. *) - Mod_subst.constant_of_delta resolver_param - (Mod_subst.constant_of_delta_kn resolver kn) + Mod_subst.constant_of_deltas_kn resolver_param resolver kn let mind_of_delta_kn kn = - let resolver,resolver_param = (delta_of_senv !global_env) in + let resolver,resolver_param = Safe_typing.delta_of_senv (safe_env ()) + in (* TODO idem *) - Mod_subst.mind_of_delta resolver_param - (Mod_subst.mind_of_delta_kn resolver kn) - -let exists_objlabel id = exists_objlabel id !global_env - -let start_library dir = - let mp,newenv = start_library dir !global_env in - global_env := newenv; - mp - -let export s = export !global_env s + Mod_subst.mind_of_deltas_kn resolver_param resolver kn -let import cenv digest = - let mp,newenv = import cenv digest !global_env in - global_env := newenv; - mp +(** Operations on libraries *) +let start_library dir = globalize (Safe_typing.start_library dir) +let export ?except s = Safe_typing.export ?except (safe_env ()) s +let import c u d = globalize (Safe_typing.import c u d) -(*s Function to get an environment from the constants part of the global +(** Function to get an environment from the constants part of the global environment and a given context. *) let env_of_context hyps = reset_with_named_context hyps (env()) -open Libnames +open Globnames -let type_of_reference env = function +(** Build a fresh instance for a given context, its associated substitution and + the instantiated constraints. *) + +let type_of_global_unsafe r = + let env = env() in + match r with | VarRef id -> Environ.named_type id env - | ConstRef c -> Typeops.type_of_constant env c + | ConstRef c -> + let cb = Environ.lookup_constant c env in + let univs = + Declareops.universes_of_polymorphic_constant + (Environ.opaque_tables env) cb in + let ty = Typeops.type_of_constant_type env cb.Declarations.const_type in + Vars.subst_instance_constr (Univ.UContext.instance univs) ty | IndRef ind -> - let specif = Inductive.lookup_mind_specif env ind in - Inductive.type_of_inductive env specif + let (mib, oib as specif) = Inductive.lookup_mind_specif env ind in + let inst = + if mib.Declarations.mind_polymorphic then + Univ.UContext.instance mib.Declarations.mind_universes + else Univ.Instance.empty + in + Inductive.type_of_inductive env (specif, inst) | ConstructRef cstr -> - let specif = - Inductive.lookup_mind_specif env (inductive_of_constructor cstr) in - Inductive.type_of_constructor cstr specif - -let type_of_global t = type_of_reference (env ()) t - + let (mib,oib as specif) = Inductive.lookup_mind_specif env (inductive_of_constructor cstr) in + let inst = Univ.UContext.instance mib.Declarations.mind_universes in + Inductive.type_of_constructor (cstr,inst) specif + +let type_of_global_in_context env r = + let open Declarations in + match r with + | VarRef id -> Environ.named_type id env, Univ.UContext.empty + | ConstRef c -> + let cb = Environ.lookup_constant c env in + let univs = + Declareops.universes_of_polymorphic_constant + (Environ.opaque_tables env) cb in + Typeops.type_of_constant_type env cb.Declarations.const_type, univs + | IndRef ind -> + let (mib, oib as specif) = Inductive.lookup_mind_specif env ind in + let univs = + if mib.mind_polymorphic then mib.mind_universes + else Univ.UContext.empty + in Inductive.type_of_inductive env (specif, Univ.UContext.instance univs), univs + | ConstructRef cstr -> + let (mib,oib as specif) = Inductive.lookup_mind_specif env (inductive_of_constructor cstr) in + let univs = + if mib.mind_polymorphic then mib.mind_universes + else Univ.UContext.empty + in + let inst = Univ.UContext.instance univs in + Inductive.type_of_constructor (cstr,inst) specif, univs + +let universes_of_global env r = + let open Declarations in + match r with + | VarRef id -> Univ.UContext.empty + | ConstRef c -> + let cb = Environ.lookup_constant c env in + Declareops.universes_of_polymorphic_constant + (Environ.opaque_tables env) cb + | IndRef ind -> + let (mib, oib) = Inductive.lookup_mind_specif env ind in + Univ.instantiate_univ_context mib.mind_universes + | ConstructRef cstr -> + let (mib,oib) = Inductive.lookup_mind_specif env (inductive_of_constructor cstr) in + Univ.instantiate_univ_context mib.mind_universes + +let universes_of_global gr = + universes_of_global (env ()) gr + +let is_polymorphic r = + let env = env() in + match r with + | VarRef id -> false + | ConstRef c -> Environ.polymorphic_constant c env + | IndRef ind -> Environ.polymorphic_ind ind env + | ConstructRef cstr -> Environ.polymorphic_ind (inductive_of_constructor cstr) env + +let is_template_polymorphic r = + let env = env() in + match r with + | VarRef id -> false + | ConstRef c -> Environ.template_polymorphic_constant c env + | IndRef ind -> Environ.template_polymorphic_ind ind env + | ConstructRef cstr -> Environ.template_polymorphic_ind (inductive_of_constructor cstr) env + +let current_dirpath () = + Safe_typing.current_dirpath (safe_env ()) + +let with_global f = + let (a, ctx) = f (env ()) (current_dirpath ()) in + push_context_set ctx; a (* spiwack: register/unregister functions for retroknowledge *) let register field value by_clause = - let entry = kind_of_term value in - let senv = Safe_typing.register !global_env field entry by_clause in - global_env := senv + globalize0 (Safe_typing.register field value by_clause) + +let register_inline c = globalize0 (Safe_typing.register_inline c) +let set_strategy k l = + GlobalSafeEnv.set_safe_env (Safe_typing.set_strategy (safe_env ()) k l) diff --git a/library/global.mli b/library/global.mli index e1cd5c7b..af23d9b7 100644 --- a/library/global.mli +++ b/library/global.mli @@ -1,106 +1,143 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* safe_environment +val safe_env : unit -> Safe_typing.safe_environment val env : unit -> Environ.env -val env_is_empty : unit -> bool +val env_is_initial : unit -> bool -val universes : unit -> universes +val universes : unit -> Univ.universes val named_context_val : unit -> Environ.named_context_val -val named_context : unit -> Sign.named_context +val named_context : unit -> Context.named_context -val env_is_empty : unit -> bool +(** {6 Enriching the global environment } *) -(** {6 Extending env with variables and local definitions } *) -val push_named_assum : (identifier * types) -> Univ.constraints -val push_named_def : (identifier * constr * types option) -> Univ.constraints +(** Changing the (im)predicativity of the system *) +val set_engagement : Declarations.engagement -> unit +val set_type_in_type : unit -> unit -(** {6 ... } *) -(** Adding constants, inductives, modules and module types. All these - functions verify that given names match those generated by kernel *) +(** Variables, Local definitions, constants, inductive types *) + +val push_named_assum : (Id.t * Constr.types) Univ.in_universe_context_set -> unit +val push_named_def : (Id.t * Entries.definition_entry) -> unit val add_constant : - dir_path -> identifier -> global_declaration -> constant -val add_mind : - dir_path -> identifier -> mutual_inductive_entry -> mutual_inductive - -val add_module : - identifier -> module_entry -> inline -> module_path * delta_resolver -val add_modtype : - identifier -> module_struct_entry -> inline -> module_path -val add_include : - module_struct_entry -> bool -> inline -> delta_resolver + DirPath.t -> Id.t -> Safe_typing.global_declaration -> constant +val add_mind : + DirPath.t -> Id.t -> Entries.mutual_inductive_entry -> mutual_inductive -val add_constraints : constraints -> unit +(** Extra universe constraints *) +val add_constraints : Univ.constraints -> unit -val set_engagement : engagement -> unit +val push_context : Univ.universe_context -> unit +val push_context_set : Univ.universe_context_set -> unit -(** {6 Interactive modules and module types } - Both [start_*] functions take the [dir_path] argument to create a - [mod_self_id]. This should be the name of the compilation unit. *) +(** Non-interactive modules and module types *) -(** [start_*] functions return the [module_path] valid for components - of the started module / module type *) +val add_module : + Id.t -> Entries.module_entry -> Declarations.inline -> + module_path * Mod_subst.delta_resolver +val add_modtype : + Id.t -> Entries.module_type_entry -> Declarations.inline -> module_path +val add_include : + Entries.module_struct_entry -> bool -> Declarations.inline -> + Mod_subst.delta_resolver -val start_module : identifier -> module_path +(** Interactive modules and module types *) -val end_module : Summary.frozen ->identifier -> - (module_struct_entry * inline) option -> module_path * delta_resolver +val start_module : Id.t -> module_path +val start_modtype : Id.t -> module_path -val add_module_parameter : - mod_bound_id -> module_struct_entry -> inline -> delta_resolver +val end_module : Summary.frozen -> Id.t -> + (Entries.module_struct_entry * Declarations.inline) option -> + module_path * MBId.t list * Mod_subst.delta_resolver -val start_modtype : identifier -> module_path -val end_modtype : Summary.frozen -> identifier -> module_path -val pack_module : unit -> module_body +val end_modtype : Summary.frozen -> Id.t -> module_path * MBId.t list +val add_module_parameter : + MBId.t -> Entries.module_struct_entry -> Declarations.inline -> + Mod_subst.delta_resolver + +(** {6 Queries in the global environment } *) + +val lookup_named : variable -> Context.named_declaration +val lookup_constant : constant -> Declarations.constant_body +val lookup_inductive : inductive -> + Declarations.mutual_inductive_body * Declarations.one_inductive_body +val lookup_pinductive : Constr.pinductive -> + Declarations.mutual_inductive_body * Declarations.one_inductive_body +val lookup_mind : mutual_inductive -> Declarations.mutual_inductive_body +val lookup_module : module_path -> Declarations.module_body +val lookup_modtype : module_path -> Declarations.module_type_body +val exists_objlabel : Label.t -> bool -(** Queries *) -val lookup_named : variable -> named_declaration -val lookup_constant : constant -> constant_body -val lookup_inductive : inductive -> mutual_inductive_body * one_inductive_body -val lookup_mind : mutual_inductive -> mutual_inductive_body -val lookup_module : module_path -> module_body -val lookup_modtype : module_path -> module_type_body val constant_of_delta_kn : kernel_name -> constant val mind_of_delta_kn : kernel_name -> mutual_inductive -val exists_objlabel : label -> bool -(** Compiled modules *) -val start_library : dir_path -> module_path -val export : dir_path -> module_path * compiled_library -val import : compiled_library -> Digest.t -> module_path +val opaque_tables : unit -> Opaqueproof.opaquetab +val body_of_constant : constant -> Term.constr option +val body_of_constant_body : Declarations.constant_body -> Term.constr option +val constraints_of_constant_body : + Declarations.constant_body -> Univ.constraints +val universes_of_constant_body : + Declarations.constant_body -> Univ.universe_context + +(** {6 Compiled libraries } *) + +val start_library : DirPath.t -> module_path +val export : ?except:Future.UUIDSet.t -> DirPath.t -> + module_path * Safe_typing.compiled_library * Safe_typing.native_library +val import : + Safe_typing.compiled_library -> Univ.universe_context_set -> Safe_typing.vodigest -> + module_path * Nativecode.symbol array + +(** {6 Misc } *) -(** {6 ... } *) (** Function to get an environment from the constants part of the global * environment and a given context. *) -val type_of_global : Libnames.global_reference -> types val env_of_context : Environ.named_context_val -> Environ.env +val join_safe_environment : ?except:Future.UUIDSet.t -> unit -> unit + +val is_polymorphic : Globnames.global_reference -> bool +val is_template_polymorphic : Globnames.global_reference -> bool + +val type_of_global_in_context : Environ.env -> + Globnames.global_reference -> Constr.types Univ.in_universe_context +val type_of_global_unsafe : Globnames.global_reference -> Constr.types + +(** Returns the universe context of the global reference (whatever it's polymorphic status is). *) +val universes_of_global : Globnames.global_reference -> Univ.universe_context + +(** {6 Retroknowledge } *) + +val register : + Retroknowledge.field -> Term.constr -> Term.constr -> unit + +val register_inline : constant -> unit + +(** {6 Oracle } *) + +val set_strategy : Names.constant Names.tableKey -> Conv_oracle.level -> unit + +(* Modifies the global state, registering new universes *) + +val current_dirpath : unit -> Names.dir_path + +val with_global : (Environ.env -> Names.dir_path -> 'a Univ.in_universe_context_set) -> 'a -(** spiwack: register/unregister function for retroknowledge *) -val register : Retroknowledge.field -> constr -> constr -> unit +val global_env_summary_name : string diff --git a/library/globnames.ml b/library/globnames.ml new file mode 100644 index 00000000..5eb091af --- /dev/null +++ b/library/globnames.ml @@ -0,0 +1,247 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* true | _ -> false +let isConstRef = function ConstRef _ -> true | _ -> false +let isIndRef = function IndRef _ -> true | _ -> false +let isConstructRef = function ConstructRef _ -> true | _ -> false + +let eq_gr gr1 gr2 = + gr1 == gr2 || match gr1,gr2 with + | ConstRef con1, ConstRef con2 -> eq_constant con1 con2 + | IndRef kn1, IndRef kn2 -> eq_ind kn1 kn2 + | ConstructRef kn1, ConstructRef kn2 -> eq_constructor kn1 kn2 + | VarRef v1, VarRef v2 -> Id.equal v1 v2 + | _ -> false + +let destVarRef = function VarRef ind -> ind | _ -> failwith "destVarRef" +let destConstRef = function ConstRef ind -> ind | _ -> failwith "destConstRef" +let destIndRef = function IndRef ind -> ind | _ -> failwith "destIndRef" +let destConstructRef = function ConstructRef ind -> ind | _ -> failwith "destConstructRef" + +let subst_constructor subst (ind,j as ref) = + let ind' = subst_ind subst ind in + if ind==ind' then ref, mkConstruct ref + else (ind',j), mkConstruct (ind',j) + +let subst_global_reference subst ref = match ref with + | VarRef var -> ref + | ConstRef kn -> + let kn' = subst_constant subst kn in + if kn==kn' then ref else ConstRef kn' + | IndRef ind -> + let ind' = subst_ind subst ind in + if ind==ind' then ref else IndRef ind' + | ConstructRef ((kn,i),j as c) -> + let c',t = subst_constructor subst c in + if c'==c then ref else ConstructRef c' + +let subst_global subst ref = match ref with + | VarRef var -> ref, mkVar var + | ConstRef kn -> + let kn',t = subst_con_kn subst kn in + if kn==kn' then ref, mkConst kn else ConstRef kn', t + | IndRef ind -> + let ind' = subst_ind subst ind in + if ind==ind' then ref, mkInd ind else IndRef ind', mkInd ind' + | ConstructRef ((kn,i),j as c) -> + let c',t = subst_constructor subst c in + if c'==c then ref,t else ConstructRef c', t + +let canonical_gr = function + | ConstRef con -> ConstRef(constant_of_kn(canonical_con con)) + | IndRef (kn,i) -> IndRef(mind_of_kn(canonical_mind kn),i) + | ConstructRef ((kn,i),j )-> ConstructRef((mind_of_kn(canonical_mind kn),i),j) + | VarRef id -> VarRef id + +let global_of_constr c = match kind_of_term c with + | Const (sp,u) -> ConstRef sp + | Ind (ind_sp,u) -> IndRef ind_sp + | Construct (cstr_cp,u) -> ConstructRef cstr_cp + | Var id -> VarRef id + | _ -> raise Not_found + +let is_global c t = + match c, kind_of_term t with + | ConstRef c, Const (c', _) -> eq_constant c c' + | IndRef i, Ind (i', _) -> eq_ind i i' + | ConstructRef i, Construct (i', _) -> eq_constructor i i' + | VarRef id, Var id' -> id_eq id id' + | _ -> false + +let printable_constr_of_global = function + | VarRef id -> mkVar id + | ConstRef sp -> mkConst sp + | ConstructRef sp -> mkConstruct sp + | IndRef sp -> mkInd sp + +let reference_of_constr = global_of_constr + +let global_eq_gen eq_cst eq_ind eq_cons x y = + x == y || + match x, y with + | ConstRef cx, ConstRef cy -> eq_cst cx cy + | IndRef indx, IndRef indy -> eq_ind indx indy + | ConstructRef consx, ConstructRef consy -> eq_cons consx consy + | VarRef v1, VarRef v2 -> Id.equal v1 v2 + | (VarRef _ | ConstRef _ | IndRef _ | ConstructRef _), _ -> false + +let global_ord_gen ord_cst ord_ind ord_cons x y = + if x == y then 0 + else match x, y with + | ConstRef cx, ConstRef cy -> ord_cst cx cy + | IndRef indx, IndRef indy -> ord_ind indx indy + | ConstructRef consx, ConstructRef consy -> ord_cons consx consy + | VarRef v1, VarRef v2 -> Id.compare v1 v2 + + | VarRef _, (ConstRef _ | IndRef _ | ConstructRef _) -> -1 + | ConstRef _, VarRef _ -> 1 + | ConstRef _, (IndRef _ | ConstructRef _) -> -1 + | IndRef _, (VarRef _ | ConstRef _) -> 1 + | IndRef _, ConstructRef _ -> -1 + | ConstructRef _, (VarRef _ | ConstRef _ | IndRef _) -> 1 + +let global_hash_gen hash_cst hash_ind hash_cons gr = + let open Hashset.Combine in + match gr with + | ConstRef c -> combinesmall 1 (hash_cst c) + | IndRef i -> combinesmall 2 (hash_ind i) + | ConstructRef c -> combinesmall 3 (hash_cons c) + | VarRef id -> combinesmall 4 (Id.hash id) + +(* By default, [global_reference] are ordered on their canonical part *) + +module RefOrdered = struct + open Constant.CanOrd + type t = global_reference + let compare gr1 gr2 = + global_ord_gen compare ind_ord constructor_ord gr1 gr2 + let equal gr1 gr2 = global_eq_gen equal eq_ind eq_constructor gr1 gr2 + let hash gr = global_hash_gen hash ind_hash constructor_hash gr +end + +module RefOrdered_env = struct + open Constant.UserOrd + type t = global_reference + let compare gr1 gr2 = + global_ord_gen compare ind_user_ord constructor_user_ord gr1 gr2 + let equal gr1 gr2 = + global_eq_gen equal eq_user_ind eq_user_constructor gr1 gr2 + let hash gr = global_hash_gen hash ind_user_hash constructor_user_hash gr +end + +module Refmap = HMap.Make(RefOrdered) +module Refset = Refmap.Set + +(* Alternative sets and maps indexed by the user part of the kernel names *) + +module Refmap_env = HMap.Make(RefOrdered_env) +module Refset_env = Refmap_env.Set + +(* Extended global references *) + +type syndef_name = kernel_name + +type extended_global_reference = + | TrueGlobal of global_reference + | SynDef of syndef_name + +(* We order [extended_global_reference] via their user part + (cf. pretty printer) *) + +module ExtRefOrdered = struct + type t = extended_global_reference + + let equal x y = + x == y || + match x, y with + | TrueGlobal rx, TrueGlobal ry -> RefOrdered_env.equal rx ry + | SynDef knx, SynDef kny -> KerName.equal knx kny + | (TrueGlobal _ | SynDef _), _ -> false + + let compare x y = + if x == y then 0 + else match x, y with + | TrueGlobal rx, TrueGlobal ry -> RefOrdered_env.compare rx ry + | SynDef knx, SynDef kny -> kn_ord knx kny + | TrueGlobal _, SynDef _ -> -1 + | SynDef _, TrueGlobal _ -> 1 + + open Hashset.Combine + + let hash = function + | TrueGlobal gr -> combinesmall 1 (RefOrdered_env.hash gr) + | SynDef kn -> combinesmall 2 (KerName.hash kn) + +end + +type global_reference_or_constr = + | IsGlobal of global_reference + | IsConstr of constr + +(** {6 Temporary function to brutally form kernel names from section paths } *) + +let encode_mind dir id = MutInd.make2 (MPfile dir) (Label.of_id id) + +let encode_con dir id = Constant.make2 (MPfile dir) (Label.of_id id) + +let check_empty_section dp = + if not (DirPath.is_empty dp) then + anomaly (Pp.str "Section part should be empty!") + +let decode_mind kn = + let rec dir_of_mp = function + | MPfile dir -> DirPath.repr dir + | MPbound mbid -> + let _,_,dp = MBId.repr mbid in + let id = MBId.to_id mbid in + id::(DirPath.repr dp) + | MPdot(mp,l) -> (Label.to_id l)::(dir_of_mp mp) + in + let mp,sec_dir,l = repr_mind kn in + check_empty_section sec_dir; + (DirPath.make (dir_of_mp mp)),Label.to_id l + +let decode_con kn = + let mp,sec_dir,l = repr_con kn in + check_empty_section sec_dir; + match mp with + | MPfile dir -> (dir,Label.to_id l) + | _ -> anomaly (Pp.str "MPfile expected!") + +(** Popping one level of section in global names. + These functions are meant to be used during discharge: + user and canonical kernel names must be equal. *) + +let pop_con con = + let (mp,dir,l) = repr_con con in + Names.make_con mp (pop_dirpath dir) l + +let pop_kn kn = + let (mp,dir,l) = repr_mind kn in + Names.make_mind mp (pop_dirpath dir) l + +let pop_global_reference = function + | ConstRef con -> ConstRef (pop_con con) + | IndRef (kn,i) -> IndRef (pop_kn kn,i) + | ConstructRef ((kn,i),j) -> ConstructRef ((pop_kn kn,i),j) + | VarRef id -> anomaly (Pp.str "VarRef not poppable") diff --git a/library/globnames.mli b/library/globnames.mli new file mode 100644 index 00000000..253c20ba --- /dev/null +++ b/library/globnames.mli @@ -0,0 +1,103 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* bool +val isConstRef : global_reference -> bool +val isIndRef : global_reference -> bool +val isConstructRef : global_reference -> bool + +val eq_gr : global_reference -> global_reference -> bool +val canonical_gr : global_reference -> global_reference + +val destVarRef : global_reference -> variable +val destConstRef : global_reference -> constant +val destIndRef : global_reference -> inductive +val destConstructRef : global_reference -> constructor + +val is_global : global_reference -> constr -> bool + +val subst_constructor : substitution -> constructor -> constructor * constr +val subst_global : substitution -> global_reference -> global_reference * constr +val subst_global_reference : substitution -> global_reference -> global_reference + +(** This constr is not safe to be typechecked, universe polymorphism is not + handled here: just use for printing *) +val printable_constr_of_global : global_reference -> constr + +(** Turn a construction denoting a global reference into a global reference; + raise [Not_found] if not a global reference *) +val global_of_constr : constr -> global_reference + +(** Obsolete synonyms for constr_of_global and global_of_constr *) +val reference_of_constr : constr -> global_reference + +module RefOrdered : sig + type t = global_reference + val compare : t -> t -> int + val equal : t -> t -> bool + val hash : t -> int +end + +module RefOrdered_env : sig + type t = global_reference + val compare : t -> t -> int + val equal : t -> t -> bool + val hash : t -> int +end + +module Refset : CSig.SetS with type elt = global_reference +module Refmap : Map.ExtS + with type key = global_reference and module Set := Refset + +module Refset_env : CSig.SetS with type elt = global_reference +module Refmap_env : Map.ExtS + with type key = global_reference and module Set := Refset_env + +(** {6 Extended global references } *) + +type syndef_name = kernel_name + +type extended_global_reference = + | TrueGlobal of global_reference + | SynDef of syndef_name + +module ExtRefOrdered : sig + type t = extended_global_reference + val compare : t -> t -> int + val equal : t -> t -> bool + val hash : t -> int +end + +type global_reference_or_constr = + | IsGlobal of global_reference + | IsConstr of constr + +(** {6 Temporary function to brutally form kernel names from section paths } *) + +val encode_mind : DirPath.t -> Id.t -> mutual_inductive +val decode_mind : mutual_inductive -> DirPath.t * Id.t +val encode_con : DirPath.t -> Id.t -> constant +val decode_con : constant -> DirPath.t * Id.t + +(** {6 Popping one level of section in global names } *) + +val pop_con : constant -> constant +val pop_kn : mutual_inductive-> mutual_inductive +val pop_global_reference : global_reference -> global_reference diff --git a/library/goptions.ml b/library/goptions.ml index d92fe262..4aea3368 100644 --- a/library/goptions.ml +++ b/library/goptions.ml @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* t -> int val table : (string * key table_of_A) list ref val encode : key -> t val subst : substitution -> t -> t @@ -62,29 +71,22 @@ module MakeTable = let nick = nickname A.key let _ = - if List.mem_assoc nick !A.table then + if String.List.mem_assoc nick !A.table then error "Sorry, this table name is already used." - module MySet = Set.Make (struct type t = A.t let compare = compare end) + module MySet = Set.Make (struct type t = A.t let compare = A.compare end) - let t = ref (MySet.empty : MySet.t) - - let _ = - if A.synchronous then - let freeze () = !t in - let unfreeze c = t := c in - let init () = t := MySet.empty in - Summary.declare_summary nick - { Summary.freeze_function = freeze; - Summary.unfreeze_function = unfreeze; - Summary.init_function = init } + let t = + if A.synchronous + then Summary.ref MySet.empty ~name:nick + else ref MySet.empty let (add_option,remove_option) = if A.synchronous then let cache_options (_,(f,p)) = match f with | GOadd -> t := MySet.add p !t | GOrmv -> t := MySet.remove p !t in - let load_options i o = if i=1 then cache_options o in + let load_options i o = if Int.equal i 1 then cache_options o in let subst_options (subst,(f,p as obj)) = let p' = A.subst subst p in if p' == p then obj else @@ -105,12 +107,12 @@ module MakeTable = (fun c -> t := MySet.remove c !t)) let print_table table_name printer table = - msg (str table_name ++ + pp (str table_name ++ (hov 0 - (if MySet.is_empty table then str "None" ++ fnl () + (if MySet.is_empty table then str " None" ++ fnl () else MySet.fold - (fun a b -> printer a ++ spc () ++ b) - table (mt ()) ++ fnl ()))) + (fun a b -> spc () ++ printer a ++ b) + table (mt ()) ++ str "." ++ fnl ()))) class table_of_A () = object @@ -119,7 +121,7 @@ module MakeTable = method mem x = let y = A.encode x in let answer = MySet.mem y !t in - msg (A.member_message y answer ++ fnl ()) + msg_info (A.member_message y answer) method print = print_table A.title A.printer !t end @@ -130,7 +132,7 @@ module MakeTable = let string_table = ref [] -let get_string_table k = List.assoc (nickname k) !string_table +let get_string_table k = String.List.assoc (nickname k) !string_table module type StringConvertArg = sig @@ -144,6 +146,7 @@ module StringConvert = functor (A : StringConvertArg) -> struct type t = string type key = string + let compare = String.compare let table = string_table let encode x = x let subst _ x = x @@ -159,11 +162,12 @@ module MakeStringTable = let ref_table = ref [] -let get_ref_table k = List.assoc (nickname k) !ref_table +let get_ref_table k = String.List.assoc (nickname k) !ref_table module type RefConvertArg = sig type t + val compare : t -> t -> int val encode : reference -> t val subst : substitution -> t -> t val printer : t -> std_ppcmds @@ -177,6 +181,7 @@ module RefConvert = functor (A : RefConvertArg) -> struct type t = A.t type key = reference + let compare = A.compare let table = ref_table let encode = A.encode let subst = A.subst @@ -201,10 +206,13 @@ type 'a option_sig = { optread : unit -> 'a; optwrite : 'a -> unit } -type option_type = bool * (unit -> option_value) -> (option_value -> unit) +module OptionOrd = +struct + type t = option_name + let compare opt1 opt2 = List.compare String.compare opt1 opt2 +end -module OptionMap = - Map.Make (struct type t = option_name let compare = compare end) +module OptionMap = Map.Make(OptionOrd) let value_tab = ref OptionMap.empty @@ -216,11 +224,10 @@ let check_key key = try let _ = get_option key in error "Sorry, this option name is already used." with Not_found -> - if List.mem_assoc (nickname key) !string_table - or List.mem_assoc (nickname key) !ref_table + if String.List.mem_assoc (nickname key) !string_table + || String.List.mem_assoc (nickname key) !ref_table then error "Sorry, this option name is already used." -open Summary open Libobject open Lib @@ -251,10 +258,10 @@ let declare_option cast uncast discharge_function = (fun (_,v) -> Some v); load_function = (fun _ (_,v) -> write v)} in - let _ = declare_summary (nickname key) - { freeze_function = read; - unfreeze_function = write; - init_function = (fun () -> write default) } + let _ = Summary.declare_summary (nickname key) + { Summary.freeze_function = (fun _ -> read ()); + Summary.unfreeze_function = write; + Summary.init_function = (fun () -> write default) } in begin fun v -> add_anonymous_leaf (decl_obj v) end , begin fun v -> add_anonymous_leaf (ldecl_obj v) end , @@ -273,15 +280,15 @@ type 'a write_function = 'a -> unit let declare_int_option = declare_option (fun v -> IntValue v) - (function IntValue v -> v | _ -> anomaly "async_option") + (function IntValue v -> v | _ -> anomaly (Pp.str "async_option")) let declare_bool_option = declare_option (fun v -> BoolValue v) - (function BoolValue v -> v | _ -> anomaly "async_option") + (function BoolValue v -> v | _ -> anomaly (Pp.str "async_option")) let declare_string_option = declare_option (fun v -> StringValue v) - (function StringValue v -> v | _ -> anomaly "async_option") + (function StringValue v -> v | _ -> anomaly (Pp.str "async_option")) (* 3- User accessible commands *) @@ -326,12 +333,12 @@ let set_int_option_value_gen locality = set_option_value locality check_int_value let set_bool_option_value_gen locality key v = try set_option_value locality check_bool_value key v - with UserError (_,s) -> Flags.if_warn msg_warning s + with UserError (_,s) -> msg_warning s let set_string_option_value_gen locality = set_option_value locality check_string_value let unset_option_value_gen locality key = try set_option_value locality check_unset_value key () - with UserError (_,s) -> Flags.if_warn msg_warning s + with UserError (_,s) -> msg_warning s let set_int_option_value = set_int_option_value_gen None let set_bool_option_value = set_bool_option_value_gen None @@ -346,19 +353,16 @@ let msg_option_value (name,v) = | IntValue (Some n) -> int n | IntValue None -> str "undefined" | StringValue s -> str s -(* | IdentValue r -> pr_global_env Idset.empty r *) +(* | IdentValue r -> pr_global_env Id.Set.empty r *) let print_option_value key = let (name, depr, (_,read,_,_,_)) = get_option key in let s = read () in match s with | BoolValue b -> - msg (str ("The "^name^" mode is "^(if b then "on" else "off")) ++ - fnl ()) + msg_info (str ("The "^name^" mode is "^(if b then "on" else "off"))) | _ -> - msg (str ("Current value of "^name^" is ") ++ - msg_option_value (name,s) ++ fnl ()) - + msg_info (str ("Current value of "^name^" is ") ++ msg_option_value (name, s)) let get_tables () = let tables = !value_tab in @@ -379,27 +383,26 @@ let print_tables () = if depr then msg ++ str " [DEPRECATED]" ++ fnl () else msg ++ fnl () in - msg - (str "Synchronous options:" ++ fnl () ++ - OptionMap.fold - (fun key (name, depr, (sync,read,_,_,_)) p -> - if sync then p ++ print_option key name (read ()) depr - else p) - !value_tab (mt ()) ++ - str "Asynchronous options:" ++ fnl () ++ - OptionMap.fold - (fun key (name, depr, (sync,read,_,_,_)) p -> - if sync then p - else p ++ print_option key name (read ()) depr) - !value_tab (mt ()) ++ - str "Tables:" ++ fnl () ++ - List.fold_right - (fun (nickkey,_) p -> p ++ str (" "^nickkey) ++ fnl ()) - !string_table (mt ()) ++ - List.fold_right - (fun (nickkey,_) p -> p ++ str (" "^nickkey) ++ fnl ()) - !ref_table (mt ()) ++ - fnl () - ) + str "Synchronous options:" ++ fnl () ++ + OptionMap.fold + (fun key (name, depr, (sync,read,_,_,_)) p -> + if sync then p ++ print_option key name (read ()) depr + else p) + !value_tab (mt ()) ++ + str "Asynchronous options:" ++ fnl () ++ + OptionMap.fold + (fun key (name, depr, (sync,read,_,_,_)) p -> + if sync then p + else p ++ print_option key name (read ()) depr) + !value_tab (mt ()) ++ + str "Tables:" ++ fnl () ++ + List.fold_right + (fun (nickkey,_) p -> p ++ str (" "^nickkey) ++ fnl ()) + !string_table (mt ()) ++ + List.fold_right + (fun (nickkey,_) p -> p ++ str (" "^nickkey) ++ fnl ()) + !ref_table (mt ()) ++ + fnl () + diff --git a/library/goptions.mli b/library/goptions.mli index 1b51a7f7..1c44f890 100644 --- a/library/goptions.mli +++ b/library/goptions.mli @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* t -> int val encode : reference -> t val subst : substitution -> t -> t val printer : t -> std_ppcmds @@ -164,7 +161,20 @@ val set_string_option_value : option_name -> string -> unit val print_option_value : option_name -> unit -val get_tables : unit -> Goptionstyp.option_state OptionMap.t -val print_tables : unit -> unit +type option_value = + | BoolValue of bool + | IntValue of int option + | StringValue of string + +(** Summary of an option status *) +type option_state = { + opt_sync : bool; + opt_depr : bool; + opt_name : string; + opt_value : option_value; +} + +val get_tables : unit -> option_state OptionMap.t +val print_tables : unit -> std_ppcmds val error_undeclared_key : option_name -> 'a diff --git a/library/goptionstyp.mli b/library/goptionstyp.mli deleted file mode 100644 index f23055b5..00000000 --- a/library/goptionstyp.mli +++ /dev/null @@ -1,26 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* - EvalConstRef(constant_of_kn(canonical_con con)) - | k -> k - in - Pervasives.compare (make_name x) (make_name y) + let compare gr1 gr2 = match gr1, gr2 with + | EvalVarRef id1, EvalVarRef id2 -> Id.compare id1 id2 + | EvalVarRef _, EvalConstRef _ -> -1 + | EvalConstRef c1, EvalConstRef c2 -> + Constant.CanOrd.compare c1 c2 + | EvalConstRef _, EvalVarRef _ -> 1 end module Evalrefmap = Map.Make (Evalreford) -let head_map = ref Evalrefmap.empty - -let init () = head_map := Evalrefmap.empty - -let freeze () = !head_map - -let unfreeze hm = head_map := hm - -let _ = - Summary.declare_summary "Head_decl" - { Summary.freeze_function = freeze; - Summary.unfreeze_function = unfreeze; - Summary.init_function = init } +let head_map = Summary.ref Evalrefmap.empty ~name:"Head_decl" let variable_head id = Evalrefmap.find (EvalVarRef id) !head_map let constant_head cst = Evalrefmap.find (EvalConstRef cst) !head_map let kind_of_head env t = - let rec aux k l t b = match kind_of_term (Reduction.whd_betaiotazeta t) with + let rec aux k l t b = match kind_of_term (Reduction.whd_betaiotazeta env t) with | Rel n when n > k -> NotImmediatelyComputableHead | Rel n -> FlexibleHead (k,k+1-n,List.length l,b) | Var id -> @@ -80,18 +66,27 @@ let kind_of_head env t = match pi2 (lookup_named id env) with | Some c -> aux k l c b | None -> NotImmediatelyComputableHead) - | Const cst -> + | Const (cst,_) -> (try on_subterm k l b (constant_head cst) with Not_found -> assert false) | Construct _ | CoFix _ -> if b then NotImmediatelyComputableHead else ConstructorHead | Sort _ | Ind _ | Prod _ -> RigidHead RigidType | Cast (c,_,_) -> aux k l c b - | Lambda (_,_,c) when l = [] -> assert (not b); aux (k+1) [] c b - | Lambda (_,_,c) -> aux k (List.tl l) (subst1 (List.hd l) c) b + | Lambda (_,_,c) -> + begin match l with + | [] -> + let () = assert (not b) in + aux (k + 1) [] c b + | h :: l -> aux k l (subst1 h c) b + end | LetIn _ -> assert false | Meta _ | Evar _ -> NotImmediatelyComputableHead | App (c,al) -> aux k (Array.to_list al @ l) c b + | Proj (p,c) -> + (try on_subterm k (c :: l) b (constant_head (Projection.constant p)) + with Not_found -> assert false) + | Case (_,_,c,_) -> aux k [] c true | Fix ((i,j),_) -> let n = i.(j) in @@ -113,18 +108,26 @@ let kind_of_head env t = k+n-m,[],a else (* enough arguments to [cst] *) - k,list_skipn n l,List.nth l (i-1) in - let l' = list_tabulate (fun _ -> mkMeta 0) q @ rest in - aux k' l' a (with_subcase or with_case) + k,List.skipn n l,List.nth l (i-1) in + let l' = List.make q (mkMeta 0) @ rest in + aux k' l' a (with_subcase || with_case) | ConstructorHead when with_case -> NotImmediatelyComputableHead | x -> x in aux 0 [] t false +(* FIXME: maybe change interface here *) let compute_head = function | EvalConstRef cst -> - (match constant_opt_value (Global.env()) cst with + let env = Global.env() in + let cb = Environ.lookup_constant cst env in + let is_Def = function Declarations.Def _ -> true | _ -> false in + let body = + if cb.Declarations.const_proj = None && is_Def cb.Declarations.const_body + then Declareops.body_of_constant (Environ.opaque_tables env) cb else None + in + (match body with | None -> RigidHead (RigidParameter cst) - | Some c -> kind_of_head (Global.env()) c) + | Some c -> kind_of_head env c) | EvalVarRef id -> (match pi2 (Global.lookup_named id) with | Some c when not (Decls.variable_opacity id) -> @@ -147,8 +150,8 @@ let cache_head o = let subst_head_approximation subst = function | RigidHead (RigidParameter cst) as k -> - let cst,c = subst_con subst cst in - if isConst c && eq_constant (destConst c) cst then + let cst,c = subst_con_kn subst cst in + if isConst c && eq_constant (fst (destConst c)) cst then (* A change of the prefix of the constant *) k else @@ -181,15 +184,3 @@ let inHead : head_obj -> obj = let declare_head c = let hd = compute_head c in add_anonymous_leaf (inHead (c,hd)) - -(** Printing *) - -let pr_head = function -| RigidHead (RigidParameter cst) -> str "rigid constant " ++ pr_con cst -| RigidHead (RigidType) -> str "rigid type" -| RigidHead (RigidVar id) -> str "rigid variable " ++ pr_id id -| ConstructorHead -> str "constructor" -| FlexibleHead (k,n,p,b) -> int n ++ str "th of " ++ int k ++ str " binders applied to " ++ int p ++ str " arguments" ++ (if b then str " (with case)" else mt()) -| NotImmediatelyComputableHead -> str "unknown" - - diff --git a/library/heads.mli b/library/heads.mli index 9ec49648..52f43824 100644 --- a/library/heads.mli +++ b/library/heads.mli @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* begin - implicit_args := oflags; - raise reraise - end + with reraise -> + let reraise = Errors.push reraise in + let () = implicit_args := oflags in + iraise reraise let set_maximality imps b = (* Force maximal insertion on ending implicits (compatibility) *) - b || List.for_all ((<>) None) imps + let is_set x = match x with None -> false | _ -> true in + b || List.for_all is_set imps (*s Computation of implicit arguments *) @@ -112,6 +113,18 @@ type argument_position = | Conclusion | Hyp of int +let argument_position_eq p1 p2 = match p1, p2 with +| Conclusion, Conclusion -> true +| Hyp h1, Hyp h2 -> Int.equal h1 h2 +| _ -> false + +let explicitation_eq ex1 ex2 = match ex1, ex2 with +| ExplByPos (i1, id1), ExplByPos (i2, id2) -> + Int.equal i1 i2 && Option.equal Id.equal id1 id2 +| ExplByName id1, ExplByName id2 -> + Id.equal id1 id2 +| _ -> false + type implicit_explanation = | DepRigid of argument_position | DepFlex of argument_position @@ -131,10 +144,10 @@ let update pos rig (na,st) = | Some (DepRigid n as x) -> if argument_less (pos,n) then DepRigid pos else x | Some (DepFlexAndRigid (fpos,rpos) as x) -> - if argument_less (pos,fpos) or pos=fpos then DepRigid pos else + if argument_less (pos,fpos) || argument_position_eq pos fpos then DepRigid pos else if argument_less (pos,rpos) then DepFlexAndRigid (fpos,pos) else x | Some (DepFlex fpos) -> - if argument_less (pos,fpos) or pos=fpos then DepRigid pos + if argument_less (pos,fpos) || argument_position_eq pos fpos then DepRigid pos else DepFlexAndRigid (fpos,pos) | Some Manual -> assert false else @@ -155,20 +168,21 @@ let is_flexible_reference env bound depth f = | Rel n when n >= bound+depth -> (* inductive type *) false | Rel n when n >= depth -> (* previous argument *) true | Rel n -> (* since local definitions have been expanded *) false - | Const kn -> + | Const (kn,_) -> let cb = Environ.lookup_constant kn env in (match cb.const_body with Def _ -> true | _ -> false) | Var id -> - let (_,value,_) = Environ.lookup_named id env in value <> None + let (_, value, _) = Environ.lookup_named id env in + begin match value with None -> false | _ -> true end | Ind _ | Construct _ -> false | _ -> true let push_lift d (e,n) = (push_rel d e,n+1) let is_reversible_pattern bound depth f l = - isRel f & let n = destRel f in (n < bound+depth) & (n >= depth) & - array_for_all (fun c -> isRel c & destRel c < depth) l & - array_distinct l + isRel f && let n = destRel f in (n < bound+depth) && (n >= depth) && + Array.for_all (fun c -> isRel c && destRel c < depth) l && + Array.distinct l (* Precondition: rels in env are for inductive types only *) let add_free_rels_until strict strongly_strict revpat bound env m pos acc = @@ -176,15 +190,18 @@ let add_free_rels_until strict strongly_strict revpat bound env m pos acc = let hd = if strict then whd_betadeltaiota env c else c in let c = if strongly_strict then hd else c in match kind_of_term hd with - | Rel n when (n < bound+depth) & (n >= depth) -> + | Rel n when (n < bound+depth) && (n >= depth) -> let i = bound + depth - n - 1 in acc.(i) <- update pos rig acc.(i) - | App (f,l) when revpat & is_reversible_pattern bound depth f l -> + | App (f,l) when revpat && is_reversible_pattern bound depth f l -> let i = bound + depth - destRel f - 1 in acc.(i) <- update pos rig acc.(i) - | App (f,_) when rig & is_flexible_reference env bound depth f -> + | App (f,_) when rig && is_flexible_reference env bound depth f -> if strict then () else iter_constr_with_full_binders push_lift (frec false) ed c + | Proj (p,c) when rig -> + if strict then () else + iter_constr_with_full_binders push_lift (frec false) ed c | Case _ when rig -> if strict then () else iter_constr_with_full_binders push_lift (frec false) ed c @@ -192,12 +209,14 @@ let add_free_rels_until strict strongly_strict revpat bound env m pos acc = | _ -> iter_constr_with_full_binders push_lift (frec rig) ed c in - frec true (env,1) m; acc + let () = if not (Vars.noccur_between 1 bound m) then frec true (env,1) m in + acc let rec is_rigid_head t = match kind_of_term t with | Rel _ | Evar _ -> false | Ind _ | Const _ | Var _ | Sort _ -> true | Case (_,_,f,_) -> is_rigid_head f + | Proj (p,c) -> true | App (f,args) -> (match kind_of_term f with | Fix ((fi,i),_) -> is_rigid_head (args.(fi.(i))) @@ -238,7 +257,7 @@ let compute_implicits_gen strict strongly_strict revpat contextual all env t = let compute_implicits_flags env f all t = compute_implicits_gen - (f.strict or f.strongly_strict) f.strongly_strict + (f.strict || f.strongly_strict) f.strongly_strict f.reversible_pattern f.contextual all env t let compute_auto_implicits env flags enriching t = @@ -256,7 +275,7 @@ type force_inference = bool (* true = always infer, never turn into evar/subgoal type implicit_status = (* None = Not implicit *) - (identifier * implicit_explanation * (maximal_insertion * force_inference)) option + (Id.t * implicit_explanation * (maximal_insertion * force_inference)) option type implicit_side_condition = DefaultImpArgs | LessArgsThan of int @@ -267,23 +286,23 @@ let is_status_implicit = function | _ -> true let name_of_implicit = function - | None -> anomaly "Not an implicit argument" + | None -> anomaly (Pp.str "Not an implicit argument") | Some (id,_,_) -> id let maximal_insertion_of = function | Some (_,_,(b,_)) -> b - | None -> anomaly "Not an implicit argument" + | None -> anomaly (Pp.str "Not an implicit argument") let force_inference_of = function | Some (_, _, (_, b)) -> b - | None -> anomaly "Not an implicit argument" + | None -> anomaly (Pp.str "Not an implicit argument") (* [in_ctx] means we know the expected type, [n] is the index of the argument *) let is_inferable_implicit in_ctx n = function | None -> false - | Some (_,DepRigid (Hyp p),_) -> in_ctx or n >= p + | Some (_,DepRigid (Hyp p),_) -> in_ctx || n >= p | Some (_,DepFlex (Hyp p),_) -> false - | Some (_,DepFlexAndRigid (_,Hyp q),_) -> in_ctx or n >= q + | Some (_,DepFlexAndRigid (_,Hyp q),_) -> in_ctx || n >= q | Some (_,DepRigid Conclusion,_) -> in_ctx | Some (_,DepFlex Conclusion,_) -> false | Some (_,DepFlexAndRigid (_,Conclusion),_) -> in_ctx @@ -300,7 +319,7 @@ let positions_of_implicits (_,impls) = let rec prepare_implicits f = function | [] -> [] - | (Anonymous, Some _)::_ -> anomaly "Unnamed implicit" + | (Anonymous, Some _)::_ -> anomaly (Pp.str "Unnamed implicit") | (Name id, Some imp)::imps -> let imps' = prepare_implicits f imps in Some (id,imp,(set_maximality imps' f.maximal,true)) :: imps' @@ -310,7 +329,7 @@ let set_implicit id imp insmax = (id,(match imp with None -> Manual | Some imp -> imp),insmax) let rec assoc_by_pos k = function - (ExplByPos (k', x), b) :: tl when k = k' -> (x,b), tl + (ExplByPos (k', x), b) :: tl when Int.equal k k' -> (x,b), tl | hd :: tl -> let (x, tl) = assoc_by_pos k tl in x, hd :: tl | [] -> raise Not_found @@ -318,9 +337,9 @@ let check_correct_manual_implicits autoimps l = List.iter (function | ExplByName id,(b,fi,forced) -> if not forced then - error ("Wrong or non-dependent implicit argument name: "^(string_of_id id)^".") + error ("Wrong or non-dependent implicit argument name: "^(Id.to_string id)^".") | ExplByPos (i,_id),_t -> - if i<1 or i>List.length autoimps then + if i<1 || i>List.length autoimps then error ("Bad implicit argument number: "^(string_of_int i)^".") else errorlabstrm "" @@ -332,34 +351,41 @@ let set_manual_implicits env flags enriching autoimps l = try let (id, (b, fi, fo)), l' = assoc_by_pos k l in if fo then - let id = match id with Some id -> id | None -> id_of_string ("arg_" ^ string_of_int k) in + let id = match id with Some id -> id | None -> Id.of_string ("arg_" ^ string_of_int k) in l', Some (id,Manual,(b,fi)) else l, None with Not_found -> l, None in - if not (list_distinct l) then + if not (List.distinct l) then error ("Some parameters are referred more than once."); (* Compare with automatic implicits to recover printing data and names *) let rec merge k l = function | (Name id,imp)::imps -> let l',imp,m = try - let (b, fi, fo) = List.assoc (ExplByName id) l in - List.remove_assoc (ExplByName id) l, (Some Manual), (Some (b, fi)) + let eq = explicitation_eq in + let (b, fi, fo) = List.assoc_f eq (ExplByName id) l in + List.remove_assoc_f eq (ExplByName id) l, (Some Manual), (Some (b, fi)) with Not_found -> try let (id, (b, fi, fo)), l' = assoc_by_pos k l in l', (Some Manual), (Some (b,fi)) with Not_found -> - l,imp, if enriching && imp <> None then Some (flags.maximal,true) else None + let m = match enriching, imp with + | true, Some _ -> Some (flags.maximal, true) + | _ -> None + in + l, imp, m in let imps' = merge (k+1) l' imps in - let m = Option.map (fun (b,f) -> set_maximality imps' b, f) m in + let m = Option.map (fun (b,f) -> + (* match imp with Some Manual -> (b,f) *) + (* | _ -> *)set_maximality imps' b, f) m in Option.map (set_implicit id imp) m :: imps' | (Anonymous,imp)::imps -> let l', forced = try_forced k l in forced :: merge (k+1) l' imps - | [] when l = [] -> [] + | [] when begin match l with [] -> true | _ -> false end -> [] | [] -> check_correct_manual_implicits autoimps l; [] @@ -376,13 +402,14 @@ let compute_semi_auto_implicits env f manual t = let _,autoimpls = compute_auto_implicits env f f.auto t in [DefaultImpArgs, set_manual_implicits env f f.auto autoimpls manual] -let compute_implicits env t = compute_semi_auto_implicits env !implicit_args [] t - (*s Constants. *) let compute_constant_implicits flags manual cst = let env = Global.env () in - compute_semi_auto_implicits env flags manual (Typeops.type_of_constant env cst) + let cb = Environ.lookup_constant cst env in + let ty = Typeops.type_of_constant_type env cb.const_type in + let impls = compute_semi_auto_implicits env flags manual ty in + impls (*s Inductives and constructors. Their implicit arguments are stored in an array, indexed by the inductive number, of pairs $(i,v)$ where @@ -394,14 +421,15 @@ let compute_mib_implicits flags manual kn = let mib = lookup_mind kn env in let ar = Array.to_list - (Array.map (* No need to lift, arities contain no de Bruijn *) - (fun mip -> - (Name mip.mind_typename, None, type_of_inductive env (mib,mip))) + (Array.mapi (* No need to lift, arities contain no de Bruijn *) + (fun i mip -> + (** No need to care about constraints here *) + (Name mip.mind_typename, None, Global.type_of_global_unsafe (IndRef (kn,i)))) mib.mind_packets) in let env_ar = push_rel_context ar env in let imps_one_inductive i mip = let ind = (kn,i) in - let ar = type_of_inductive env (mib,mip) in + let ar = Global.type_of_global_unsafe (IndRef ind) in ((IndRef ind,compute_semi_auto_implicits env flags manual ar), Array.mapi (fun j c -> (ConstructRef (ind,j+1),compute_semi_auto_implicits env_ar flags manual c)) @@ -412,7 +440,7 @@ let compute_mib_implicits flags manual kn = let compute_all_mib_implicits flags manual kn = let imps = compute_mib_implicits flags manual kn in List.flatten - (array_map_to_list (fun (ind,cstrs) -> ind::Array.to_list cstrs) imps) + (Array.map_to_list (fun (ind,cstrs) -> ind::Array.to_list cstrs) imps) (*s Variables. *) @@ -434,7 +462,7 @@ let compute_global_implicits flags manual = function (* Merge a manual explicitation with an implicit_status list *) let merge_impls (cond,oldimpls) (_,newimpls) = - let oldimpls,usersuffiximpls = list_chop (List.length newimpls) oldimpls in + let oldimpls,usersuffiximpls = List.chop (List.length newimpls) oldimpls in cond, (List.map2 (fun orig ni -> match orig with | Some (_, Manual, _) -> orig @@ -453,7 +481,7 @@ type implicit_discharge_request = | ImplInteractive of global_reference * implicits_flags * implicit_interactive_request -let implicits_table = ref Refmap.empty +let implicits_table = Summary.ref Refmap.empty ~name:"implicits" let implicits_of_global ref = try @@ -466,7 +494,7 @@ let implicits_of_global ref = List.map2 (fun (t, il) rl -> t, List.map2 rename il rl) l rename_l with Not_found -> l | Invalid_argument _ -> - anomaly "renamings list and implicits list have different lenghts" + anomaly (Pp.str "renamings list and implicits list have different lenghts") with Not_found -> [DefaultImpArgs,[]] let cache_implicits_decl (ref,imps) = @@ -481,16 +509,23 @@ let subst_implicits_decl subst (r,imps as o) = let r' = fst (subst_global subst r) in if r==r' then o else (r',imps) let subst_implicits (subst,(req,l)) = - (ImplLocal,list_smartmap (subst_implicits_decl subst) l) + (ImplLocal,List.smartmap (subst_implicits_decl subst) l) let impls_of_context ctx = - List.rev_map (fun (id,impl,_,_) -> if impl = Lib.Implicit then Some (id, Manual, (true,true)) else None) - (List.filter (fun (_,_,b,_) -> b = None) ctx) + let map (id, impl, _, _) = match impl with + | Implicit -> Some (id, Manual, (true, true)) + | _ -> None + in + let is_set (_, _, b, _) = match b with + | None -> true + | Some _ -> false + in + List.rev_map map (List.filter is_set ctx) let section_segment_of_reference = function - | ConstRef con -> section_segment_of_constant con + | ConstRef con -> pi1 (section_segment_of_constant con) | IndRef (kn,_) | ConstructRef ((kn,_),_) -> - section_segment_of_mutual_inductive kn + pi1 (section_segment_of_mutual_inductive kn) | _ -> [] let adjust_side_condition p = function @@ -515,9 +550,10 @@ let discharge_implicits (_,(req,l)) = | ImplConstant (con,flags) -> (try let con' = pop_con con in - let vars = section_segment_of_constant con in + let vars,_,_ = section_segment_of_constant con in let extra_impls = impls_of_context vars in - let l' = [ConstRef con',List.map (add_section_impls vars extra_impls) (snd (List.hd l))] in + let newimpls = List.map (add_section_impls vars extra_impls) (snd (List.hd l)) in + let l' = [ConstRef con',newimpls] in Some (ImplConstant (con',flags),l') with Not_found -> (* con not defined in this section *) Some (req,l)) | ImplMutualInductive (kn,flags) -> @@ -560,13 +596,14 @@ let rebuild_implicits (req,l) = if flags.auto then let newimpls = List.hd (compute_global_implicits flags [] ref) in let p = List.length (snd newimpls) - userimplsize in - let newimpls = on_snd (list_firstn p) newimpls in + let newimpls = on_snd (List.firstn p) newimpls in [ref,List.map (fun o -> merge_impls o newimpls) oldimpls] else [ref,oldimpls] -let classify_implicits (req,_ as obj) = - if req = ImplLocal then Dispose else Substitute obj +let classify_implicits (req,_ as obj) = match req with +| ImplLocal -> Dispose +| _ -> Substitute obj type implicits_obj = implicit_discharge_request * @@ -603,14 +640,14 @@ let declare_constant_implicits con = let declare_mib_implicits kn = let flags = !implicit_args in - let imps = array_map_to_list + let imps = Array.map_to_list (fun (ind,cstrs) -> ind::(Array.to_list cstrs)) (compute_mib_implicits flags [] kn) in add_anonymous_leaf (inImplicits (ImplMutualInductive (kn,flags),List.flatten imps)) (* Declare manual implicits *) -type manual_explicitation = Topconstr.explicitation * (bool * bool * bool) +type manual_explicitation = Constrexpr.explicitation * (bool * bool * bool) type manual_implicits = manual_explicitation list @@ -632,10 +669,14 @@ let check_rigidity isrigid = if not isrigid then errorlabstrm "" (strbrk "Multiple sequences of implicit arguments available only for references that cannot be applied to an arbitrarily large number of arguments.") +let projection_implicits env p impls = + let pb = Environ.lookup_projection p env in + CList.skipn_at_least pb.Declarations.proj_npars impls + let declare_manual_implicits local ref ?enriching l = let flags = !implicit_args in let env = Global.env () in - let t = Global.type_of_global ref in + let t = Global.type_of_global_unsafe ref in let enriching = Option.default flags.auto enriching in let isrigid,autoimpls = compute_auto_implicits env flags enriching t in let l' = match l with @@ -645,7 +686,7 @@ let declare_manual_implicits local ref ?enriching l = | _ -> check_rigidity isrigid; let l = List.map (fun imps -> (imps,List.length imps)) l in - let l = Sort.list (fun (_,n1) (_,n2) -> n1 > n2) l in + let l = List.sort (fun (_,n1) (_,n2) -> n2 - n1) l in check_inclusion l; let nargs = List.length autoimpls in List.map (fun (imps,n) -> @@ -658,8 +699,9 @@ let declare_manual_implicits local ref ?enriching l = add_anonymous_leaf (inImplicits (req,[ref,l'])) let maybe_declare_manual_implicits local ref ?enriching l = - if l = [] then () - else declare_manual_implicits local ref ?enriching [l] + match l with + | [] -> () + | _ -> declare_manual_implicits local ref ?enriching [l] let extract_impargs_data impls = let rec aux p = function @@ -677,7 +719,7 @@ let lift_implicits n = let make_implicits_list l = [DefaultImpArgs, l] let rec drop_first_implicits p l = - if p = 0 then l else match l with + if Int.equal p 0 then l else match l with | _,[] as x -> x | DefaultImpArgs,imp::impls -> drop_first_implicits (p-1) (DefaultImpArgs,impls) @@ -691,18 +733,6 @@ let rec select_impargs_size n = function | (LessArgsThan p, impls)::l -> if n <= p then impls else select_impargs_size n l -let rec select_stronger_impargs = function +let select_stronger_impargs = function | [] -> [] (* Tolerance for (DefaultImpArgs,[]) *) | (_,impls)::_ -> impls - -(*s Registration as global tables *) - -let init () = implicits_table := Refmap.empty -let freeze () = !implicits_table -let unfreeze t = implicits_table := t - -let _ = - Summary.declare_summary "implicits" - { Summary.freeze_function = freeze; - Summary.unfreeze_function = unfreeze; - Summary.init_function = init } diff --git a/library/impargs.mli b/library/impargs.mli index 9038ca88..1d3a73e9 100644 --- a/library/impargs.mli +++ b/library/impargs.mli @@ -1,16 +1,15 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* bool val is_inferable_implicit : bool -> int -> implicit_status -> bool -val name_of_implicit : implicit_status -> identifier +val name_of_implicit : implicit_status -> Id.t val maximal_insertion_of : implicit_status -> bool val force_inference_of : implicit_status -> bool @@ -87,7 +86,7 @@ val positions_of_implicits : implicits_list -> int list (** A [manual_explicitation] is a tuple of a positional or named explicitation with maximal insertion, force inference and force usage flags. Forcing usage makes the argument implicit even if the automatic inference considers it not inferable. *) -type manual_explicitation = Topconstr.explicitation * +type manual_explicitation = Constrexpr.explicitation * (maximal_insertion * force_inference * bool) type manual_implicits = manual_explicitation list @@ -95,7 +94,7 @@ type manual_implicits = manual_explicitation list val compute_implicits_with_manual : env -> types -> bool -> manual_implicits -> implicit_status list -val compute_implicits_names : env -> types -> name list +val compute_implicits_names : env -> types -> Name.t list (** {6 Computation of implicits (done using the global environment). } *) @@ -130,6 +129,9 @@ val make_implicits_list : implicit_status list -> implicits_list list val drop_first_implicits : int -> implicits_list -> implicits_list +val projection_implicits : env -> projection -> implicit_status list -> + implicit_status list + val select_impargs_size : int -> implicits_list list -> implicit_status list val select_stronger_impargs : implicits_list list -> implicit_status list @@ -143,3 +145,5 @@ type implicit_discharge_request = | ImplInteractive of global_reference * implicits_flags * implicit_interactive_request +val explicitation_eq : Constrexpr.explicitation -> Constrexpr.explicitation -> bool +(** Equality on [explicitation]. *) diff --git a/library/keys.ml b/library/keys.ml new file mode 100644 index 00000000..3d277476 --- /dev/null +++ b/library/keys.ml @@ -0,0 +1,170 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* 10 + RefOrdered.hash gr + | KLam -> 0 + | KLet -> 1 + | KProd -> 2 + | KSort -> 3 + | KEvar -> 4 + | KCase -> 5 + | KFix -> 6 + | KCoFix -> 7 + | KRel -> 8 + | KMeta -> 9 + + let compare gr1 gr2 = + match gr1, gr2 with + | KGlob gr1, KGlob gr2 -> RefOrdered.compare gr1 gr2 + | _, KGlob _ -> -1 + | KGlob _, _ -> 1 + | k, k' -> Int.compare (hash k) (hash k') + + let equal k1 k2 = + match k1, k2 with + | KGlob gr1, KGlob gr2 -> RefOrdered.equal gr1 gr2 + | _, KGlob _ -> false + | KGlob _, _ -> false + | k, k' -> k == k' +end + +module Keymap = HMap.Make(KeyOrdered) +module Keyset = Keymap.Set + +(* Mapping structure for references to be considered equivalent *) + +type keys = Keyset.t Keymap.t + +let keys = Summary.ref Keymap.empty ~name:"Keys_decl" + +let add_kv k v m = + try Keymap.modify k (fun k' vs -> Keyset.add v vs) m + with Not_found -> Keymap.add k (Keyset.singleton v) m + +let add_keys k v = + keys := add_kv k v (add_kv v k !keys) + +let equiv_keys k k' = + k == k' || KeyOrdered.equal k k' || + try Keyset.mem k' (Keymap.find k !keys) + with Not_found -> false + +(** Registration of keys as an object *) + +let load_keys _ (_,(ref,ref')) = + add_keys ref ref' + +let cache_keys o = + load_keys 1 o + +let subst_key subst k = + match k with + | KGlob gr -> KGlob (subst_global_reference subst gr) + | _ -> k + +let subst_keys (subst,(k,k')) = + (subst_key subst k, subst_key subst k') + +let discharge_key = function + | KGlob g when Lib.is_in_section g -> + if isVarRef g then None else Some (KGlob (pop_global_reference g)) + | x -> Some x + +let discharge_keys (_,(k,k')) = + match discharge_key k, discharge_key k' with + | Some x, Some y -> Some (x, y) + | _ -> None + +let rebuild_keys (ref,ref') = (ref, ref') + +type key_obj = key * key + +let inKeys : key_obj -> obj = + declare_object {(default_object "KEYS") with + cache_function = cache_keys; + load_function = load_keys; + subst_function = subst_keys; + classify_function = (fun x -> Substitute x); + discharge_function = discharge_keys; + rebuild_function = rebuild_keys } + +let declare_equiv_keys ref ref' = + Lib.add_anonymous_leaf (inKeys (ref,ref')) + +let constr_key c = + let open Globnames in + try + let rec aux k = + match kind_of_term k with + | Const (c, _) -> KGlob (ConstRef c) + | Ind (i, u) -> KGlob (IndRef i) + | Construct (c,u) -> KGlob (ConstructRef c) + | Var id -> KGlob (VarRef id) + | App (f, _) -> aux f + | Proj (p, _) -> KGlob (ConstRef (Names.Projection.constant p)) + | Cast (p, _, _) -> aux p + | Lambda _ -> KLam + | Prod _ -> KProd + | Case _ -> KCase + | Fix _ -> KFix + | CoFix _ -> KCoFix + | Rel _ -> KRel + | Meta _ -> raise Not_found + | Evar _ -> raise Not_found + | Sort _ -> KSort + | LetIn _ -> KLet + in Some (aux c) + with Not_found -> None + +open Pp + +let pr_key pr_global = function + | KGlob gr -> pr_global gr + | KLam -> str"Lambda" + | KLet -> str"Let" + | KProd -> str"Product" + | KSort -> str"Sort" + | KEvar -> str"Evar" + | KCase -> str"Case" + | KFix -> str"Fix" + | KCoFix -> str"CoFix" + | KRel -> str"Rel" + | KMeta -> str"Meta" + +let pr_keyset pr_global v = + prlist_with_sep spc (pr_key pr_global) (Keyset.elements v) + +let pr_mapping pr_global k v = + pr_key pr_global k ++ str" <-> " ++ pr_keyset pr_global v + +let pr_keys pr_global = + Keymap.fold (fun k v acc -> pr_mapping pr_global k v ++ fnl () ++ acc) !keys (mt()) diff --git a/library/keys.mli b/library/keys.mli new file mode 100644 index 00000000..bfbb4c58 --- /dev/null +++ b/library/keys.mli @@ -0,0 +1,23 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* key -> unit +(** Declare two keys as being equivalent. *) + +val equiv_keys : key -> key -> bool +(** Check equivalence of keys. *) + +val constr_key : Term.constr -> key option +(** Compute the head key of a term. *) + +val pr_keys : (global_reference -> Pp.std_ppcmds) -> Pp.std_ppcmds +(** Pretty-print the mapping *) diff --git a/library/kindops.ml b/library/kindops.ml new file mode 100644 index 00000000..56048647 --- /dev/null +++ b/library/kindops.ml @@ -0,0 +1,67 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* IsDefinition d + | Proof s -> IsProof s + +let string_of_theorem_kind = function + | Theorem -> "Theorem" + | Lemma -> "Lemma" + | Fact -> "Fact" + | Remark -> "Remark" + | Property -> "Property" + | Proposition -> "Proposition" + | Corollary -> "Corollary" + +let string_of_definition_kind def = + let (locality, poly, kind) = def in + let error () = Errors.anomaly (Pp.str "Internal definition kind") in + match kind with + | Definition -> + begin match locality with + | Discharge -> "Let" + | Local -> "Local Definition" + | Global -> "Definition" + end + | Example -> + begin match locality with + | Discharge -> error () + | Local -> "Local Example" + | Global -> "Example" + end + | Coercion -> + begin match locality with + | Discharge -> error () + | Local -> "Local Coercion" + | Global -> "Coercion" + end + | SubClass -> + begin match locality with + | Discharge -> error () + | Local -> "Local SubClass" + | Global -> "SubClass" + end + | CanonicalStructure -> + begin match locality with + | Discharge -> error () + | Local -> error () + | Global -> "Canonical Structure" + end + | Instance -> + begin match locality with + | Discharge -> error () + | Local -> "Instance" + | Global -> "Global Instance" + end + | (StructureComponent|Scheme|CoFixpoint|Fixpoint|IdentityCoercion|Method) -> + Errors.anomaly (Pp.str "Internal definition kind") diff --git a/library/kindops.mli b/library/kindops.mli new file mode 100644 index 00000000..cd2e39cf --- /dev/null +++ b/library/kindops.mli @@ -0,0 +1,15 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* logical_kind +val string_of_theorem_kind : theorem_kind -> string +val string_of_definition_kind : definition_kind -> string diff --git a/library/lib.ml b/library/lib.ml index f18bbac6..9977b666 100644 --- a/library/lib.ml +++ b/library/lib.ml @@ -1,17 +1,18 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* f i (make_oname prefix id, obj)) -let load_objects = iter_objects load_object -let open_objects = iter_objects open_object +let load_objects i pr = iter_objects load_object i pr +let open_objects i pr = iter_objects open_object i pr let subst_objects subst seg = let subst_one = fun (id,obj as node) -> @@ -46,7 +47,7 @@ let subst_objects subst seg = if obj' == obj then node else (id, obj') in - list_smartmap subst_one seg + List.smartmap subst_one seg (*let load_and_subst_objects i prefix subst seg = List.rev (List.fold_left (fun seg (id,obj as node) -> @@ -59,7 +60,7 @@ let classify_segment seg = let rec clean ((substl,keepl,anticipl) as acc) = function | (_,CompilingLibrary _) :: _ | [] -> acc | ((sp,kn),Leaf o) :: stk -> - let id = Names.id_of_label (Names.label kn) in + let id = Names.Label.to_id (Names.label kn) in (match classify_object o with | Dispose -> clean acc stk | Keep o' -> @@ -90,7 +91,7 @@ let segment_of_objects prefix = sections, but on the contrary there are many constructions of section paths based on the library path. *) -let initial_prefix = default_library,(Names.initial_path,Names.empty_dirpath) +let initial_prefix = default_library,(Names.initial_path,Names.DirPath.empty) let lib_stk = ref ([] : library_segment) @@ -103,15 +104,13 @@ let library_dp () = module path and relative section path *) let path_prefix = ref initial_prefix -let sections_depth () = - List.length (Names.repr_dirpath (snd (snd !path_prefix))) - -let sections_are_opened () = - match Names.repr_dirpath (snd (snd !path_prefix)) with - [] -> false - | _ -> true - let cwd () = fst !path_prefix +let current_prefix () = snd !path_prefix +let current_mp () = fst (snd !path_prefix) +let current_sections () = snd (snd !path_prefix) + +let sections_depth () = List.length (Names.DirPath.repr (current_sections ())) +let sections_are_opened () = not (Names.DirPath.is_empty (current_sections ())) let cwd_except_section () = Libnames.pop_dirpath_n (sections_depth ()) (cwd ()) @@ -122,26 +121,14 @@ let current_dirpath sec = let make_path id = Libnames.make_path (cwd ()) id -let make_path_except_section id = Libnames.make_path (cwd_except_section ()) id - -let path_of_include () = - let dir = Names.repr_dirpath (cwd ()) in - let new_dir = List.tl dir in - let id = List.hd dir in - Libnames.make_path (Names.make_dirpath new_dir) id - -let current_prefix () = snd !path_prefix +let make_path_except_section id = + Libnames.make_path (cwd_except_section ()) id let make_kn id = let mp,dir = current_prefix () in - Names.make_kn mp dir (Names.label_of_id id) - -let make_con id = - let mp,dir = current_prefix () in - Names.make_con mp dir (Names.label_of_id id) + Names.make_kn mp dir (Names.Label.of_id id) - -let make_oname id = make_path id, make_kn id +let make_oname id = Libnames.make_oname !path_prefix id let recalc_path_prefix () = let rec recalc = function @@ -155,7 +142,7 @@ let recalc_path_prefix () = let pop_path_prefix () = let dir,(mp,sec) = !path_prefix in - path_prefix := fst (split_dirpath dir), (mp, fst (split_dirpath sec)) + path_prefix := pop_dirpath dir, (mp, pop_dirpath sec) let find_entry_p p = let rec find = function @@ -164,13 +151,6 @@ let find_entry_p p = in find !lib_stk -let find_split_p p = - let rec find = function - | [] -> raise Not_found - | ent::l -> if p ent then ent,l else find l - in - find !lib_stk - let split_lib_gen test = let rec collect after equal = function | hd::before when test hd -> collect after (hd::equal) before @@ -194,16 +174,23 @@ let split_lib_gen test = | None -> error "no such entry" | Some r -> r -let split_lib sp = split_lib_gen (fun x -> fst x = sp) +let eq_object_name (fp1, kn1) (fp2, kn2) = + eq_full_path fp1 fp2 && Names.KerName.equal kn1 kn2 + +let split_lib sp = + let is_sp (nsp, _) = eq_object_name sp nsp in + split_lib_gen is_sp let split_lib_at_opening sp = - let is_sp = function - | x,(OpenedSection _|OpenedModule _|CompilingLibrary _) -> x = sp + let is_sp (nsp, obj) = match obj with + | OpenedSection _ | OpenedModule _ | CompilingLibrary _ -> + eq_object_name nsp sp | _ -> false in - let a,s,b = split_lib_gen is_sp in - assert (List.tl s = []); - (a,List.hd s,b) + let a, s, b = split_lib_gen is_sp in + match s with + | [obj] -> (a, obj, b) + | _ -> assert false (* Adding operations. *) @@ -212,16 +199,13 @@ let add_entry sp node = let anonymous_id = let n = ref 0 in - fun () -> incr n; Names.id_of_string ("_" ^ (string_of_int !n)) + fun () -> incr n; Names.Id.of_string ("_" ^ (string_of_int !n)) let add_anonymous_entry node = - let id = anonymous_id () in - let name = make_oname id in - add_entry name node; - name + add_entry (make_oname (anonymous_id ())) node let add_leaf id obj = - if fst (current_prefix ()) = Names.initial_path then + if Names.ModPath.equal (current_mp ()) Names.initial_path then error ("No session module started (use -top dir)"); let oname = make_oname id in cache_object (oname,obj); @@ -250,7 +234,8 @@ let add_anonymous_leaf obj = add_entry oname (Leaf obj) let add_frozen_state () = - let _ = add_anonymous_entry (FrozenState (freeze_summaries())) in () + add_anonymous_entry + (FrozenState (Summary.freeze_summaries ~marshallable:`No)) (* Modules. *) @@ -271,19 +256,17 @@ let current_mod_id () = let start_mod is_type export id mp fs = - let dir = add_dirpath_suffix (fst !path_prefix) id in - let prefix = dir,(mp,Names.empty_dirpath) in - let sp = make_path id in - let oname = sp, make_kn id in + let dir = add_dirpath_suffix (cwd ()) id in + let prefix = dir,(mp,Names.DirPath.empty) in let exists = - if is_type then Nametab.exists_cci sp else Nametab.exists_module dir + if is_type then Nametab.exists_cci (make_path id) + else Nametab.exists_module dir in if exists then errorlabstrm "open_module" (pr_id id ++ str " already exists"); - add_entry oname (OpenedModule (is_type,export,prefix,fs)); + add_entry (make_oname id) (OpenedModule (is_type,export,prefix,fs)); path_prefix := prefix; prefix -(* add_frozen_state () must be called in declaremods *) let start_module = start_mod false let start_modtype = start_mod true None @@ -297,7 +280,7 @@ let end_mod is_type = let oname,fs = try match find_entry_p is_opening_node with | oname,OpenedModule (ty,_,_,fs) -> - if ty = is_type then oname,fs + if ty == is_type then oname, fs else error_still_opened (module_kind ty) oname | oname,OpenedSection _ -> error_still_opened "section" oname | _ -> assert false @@ -308,31 +291,29 @@ let end_mod is_type = add_entry oname (ClosedModule (List.rev (mark::after))); let prefix = !path_prefix in recalc_path_prefix (); - (* add_frozen_state must be called after processing the module, - because we cannot recache interactive modules *) (oname, prefix, fs, after) let end_module () = end_mod false let end_modtype () = end_mod true -let contents_after = function - | None -> !lib_stk - | Some sp -> let (after,_,_) = split_lib sp in after +let contents () = !lib_stk + +let contents_after sp = let (after,_,_) = split_lib sp in after (* Modules. *) (* TODO: use check_for_module ? *) let start_compilation s mp = - if !comp_name <> None then + if !comp_name != None then error "compilation unit is already started"; - if snd (snd (!path_prefix)) <> Names.empty_dirpath then + if not (Names.DirPath.is_empty (current_sections ())) then error "some sections are already opened"; - let prefix = s, (mp, Names.empty_dirpath) in - let _ = add_anonymous_entry (CompilingLibrary prefix) in + let prefix = s, (mp, Names.DirPath.empty) in + let () = add_anonymous_entry (CompilingLibrary prefix) in comp_name := Some s; path_prefix := prefix -let end_compilation dir = +let end_compilation_checks dir = let _ = try match snd (find_entry_p is_opening_node) with | OpenedSection _ -> error "There are some open sections." @@ -347,42 +328,48 @@ let end_compilation dir = try match find_entry_p is_opening_lib with | (oname, CompilingLibrary prefix) -> oname | _ -> assert false - with Not_found -> anomaly "No module declared" + with Not_found -> anomaly (Pp.str "No module declared") in let _ = match !comp_name with - | None -> anomaly "There should be a module name..." + | None -> anomaly (Pp.str "There should be a module name...") | Some m -> - if m <> dir then anomaly - ("The current open module has name "^ (Names.string_of_dirpath m) ^ - " and not " ^ (Names.string_of_dirpath m)); + if not (Names.DirPath.equal m dir) then anomaly + (str "The current open module has name" ++ spc () ++ pr_dirpath m ++ + spc () ++ str "and not" ++ spc () ++ pr_dirpath m); in + oname + +let end_compilation oname = let (after,mark,before) = split_lib_at_opening oname in comp_name := None; !path_prefix,after (* Returns true if we are inside an opened module or module type *) -let is_module_gen which = +let is_module_gen which check = let test = function | _, OpenedModule (ty,_,_,_) -> which ty | _ -> false in try - let _ = find_entry_p test in true + match find_entry_p test with + | _, OpenedModule (ty,_,_,_) -> check ty + | _ -> assert false with Not_found -> false -let is_module_or_modtype () = is_module_gen (fun _ -> true) -let is_modtype () = is_module_gen (fun b -> b) -let is_module () = is_module_gen (fun b -> not b) +let is_module_or_modtype () = is_module_gen (fun _ -> true) (fun _ -> true) +let is_modtype () = is_module_gen (fun b -> b) (fun _ -> true) +let is_modtype_strict () = is_module_gen (fun _ -> true) (fun b -> b) +let is_module () = is_module_gen (fun b -> not b) (fun _ -> true) (* Returns the opening node of a given name *) let find_opening_node id = try let oname,entry = find_entry_p is_opening_node in let id' = basename (fst oname) in - if id <> id' then - error ("Last block to end has name "^(Names.string_of_id id')^"."); + if not (Names.Id.equal id id') then + error ("Last block to end has name "^(Names.Id.to_string id')^"."); entry with Not_found -> error "There is nothing to end." @@ -394,29 +381,39 @@ let find_opening_node id = - the list of variables on which each inductive depends in this section - the list of substitution to do at section closing *) -type binding_kind = Explicit | Implicit -type variable_info = Names.identifier * binding_kind * Term.constr option * Term.types +type variable_info = Names.Id.t * Decl_kinds.binding_kind * Term.constr option * Term.types + type variable_context = variable_info list -type abstr_list = variable_context Names.Cmap.t * variable_context Names.Mindmap.t +type abstr_info = variable_context * Univ.universe_level_subst * Univ.UContext.t + +type abstr_list = abstr_info Names.Cmap.t * abstr_info Names.Mindmap.t let sectab = - ref ([] : ((Names.identifier * binding_kind) list * Cooking.work_list * abstr_list) list) + Summary.ref ([] : ((Names.Id.t * Decl_kinds.binding_kind * + Decl_kinds.polymorphic * Univ.universe_context_set) list * + Opaqueproof.work_list * abstr_list) list) + ~name:"section-context" let add_section () = - sectab := ([],(Names.Cmap.empty,Names.Mindmap.empty),(Names.Cmap.empty,Names.Mindmap.empty)) :: !sectab + sectab := ([],(Names.Cmap.empty,Names.Mindmap.empty), + (Names.Cmap.empty,Names.Mindmap.empty)) :: !sectab -let add_section_variable id impl = +let add_section_variable id impl poly ctx = match !sectab with | [] -> () (* because (Co-)Fixpoint temporarily uses local vars *) | (vars,repl,abs)::sl -> - sectab := ((id,impl)::vars,repl,abs)::sl + sectab := ((id,impl,poly,ctx)::vars,repl,abs)::sl let extract_hyps (secs,ohyps) = let rec aux = function - | ((id,impl)::idl,(id',b,t)::hyps) when id=id' -> (id',impl,b,t) :: aux (idl,hyps) - | (id::idl,hyps) -> aux (idl,hyps) - | [], _ -> [] + | ((id,impl,poly,ctx)::idl,(id',b,t)::hyps) when Names.Id.equal id id' -> + let l, r = aux (idl,hyps) in + (id',impl,b,t) :: l, if poly then Univ.ContextSet.union r ctx else r + | ((_,_,poly,ctx)::idl,hyps) -> + let l, r = aux (idl,hyps) in + l, if poly then Univ.ContextSet.union r ctx else r + | [], _ -> [],Univ.ContextSet.empty in aux (secs,ohyps) let instance_from_variable_context sign = @@ -426,23 +423,25 @@ let instance_from_variable_context sign = | [] -> [] in Array.of_list (inst_rec sign) -let named_of_variable_context = List.map (fun (id,_,b,t) -> (id,b,t)) - +let named_of_variable_context ctx = List.map (fun (id,_,b,t) -> (id,b,t)) ctx + let add_section_replacement f g hyps = match !sectab with | [] -> () | (vars,exps,abs)::sl -> - let sechyps = extract_hyps (vars,hyps) in + let sechyps,ctx = extract_hyps (vars,hyps) in + let ctx = Univ.ContextSet.to_context ctx in + let subst, ctx = Univ.abstract_universes true ctx in let args = instance_from_variable_context (List.rev sechyps) in - sectab := (vars,f args exps,g sechyps abs)::sl + sectab := (vars,f (Univ.UContext.instance ctx,args) exps,g (sechyps,subst,ctx) abs)::sl let add_section_kn kn = let f x (l1,l2) = (l1,Names.Mindmap.add kn x l2) in - add_section_replacement f f + add_section_replacement f f -let add_section_constant kn = +let add_section_constant is_projection kn = let f x (l1,l2) = (Names.Cmap.add kn x l1,l2) in - add_section_replacement f f + add_section_replacement f f let replacement_context () = pi2 (List.hd !sectab) @@ -452,13 +451,11 @@ let section_segment_of_constant con = let section_segment_of_mutual_inductive kn = Names.Mindmap.find kn (snd (pi3 (List.hd !sectab))) -let rec list_mem_assoc x = function - | [] -> raise Not_found - | (a,_)::l -> compare a x = 0 or list_mem_assoc x l - let section_instance = function | VarRef id -> - if list_mem_assoc id (pi1 (List.hd !sectab)) then [||] + if List.exists (fun (id',_,_,_) -> Names.id_eq id id') + (pi1 (List.hd !sectab)) + then Univ.Instance.empty, [||] else raise Not_found | ConstRef con -> Names.Cmap.find con (fst (pi2 (List.hd !sectab))) @@ -468,39 +465,27 @@ let section_instance = function let is_in_section ref = try ignore (section_instance ref); true with Not_found -> false -let init_sectab () = sectab := [] -let freeze_sectab () = !sectab -let unfreeze_sectab s = sectab := s - -let _ = - Summary.declare_summary "section-context" - { Summary.freeze_function = freeze_sectab; - Summary.unfreeze_function = unfreeze_sectab; - Summary.init_function = init_sectab } +let full_replacement_context () = List.map pi2 !sectab +let full_section_segment_of_constant con = + List.map (fun (vars,_,(x,_)) -> fun hyps -> + named_of_variable_context + (try pi1 (Names.Cmap.find con x) + with Not_found -> fst (extract_hyps (vars, hyps)))) !sectab (*************) (* Sections. *) -(* XML output hooks *) -let xml_open_section = ref (fun id -> ()) -let xml_close_section = ref (fun id -> ()) - -let set_xml_open_section f = xml_open_section := f -let set_xml_close_section f = xml_close_section := f - let open_section id = let olddir,(mp,oldsec) = !path_prefix in let dir = add_dirpath_suffix olddir id in let prefix = dir, (mp, add_dirpath_suffix oldsec id) in - let name = make_path id, make_kn id (* this makes little sense however *) in if Nametab.exists_section dir then errorlabstrm "open_section" (pr_id id ++ str " already exists."); - let fs = freeze_summaries() in - add_entry name (OpenedSection (prefix, fs)); + let fs = Summary.freeze_summaries ~marshallable:`No in + add_entry (make_oname id) (OpenedSection (prefix, fs)); (*Pushed for the lifetime of the section: removed by unfrozing the summary*) Nametab.push_dir (Nametab.Until 1) dir (DirOpenSection prefix); path_prefix := prefix; - if !Flags.xml_export then !xml_open_section id; add_section () @@ -514,7 +499,7 @@ let discharge_item ((sp,_ as oname),e) = | FrozenState _ -> None | ClosedSection _ | ClosedModule _ -> None | OpenedSection _ | OpenedModule _ | CompilingLibrary _ -> - anomaly "discharge_item" + anomaly (Pp.str "discharge_item") let close_section () = let oname,fs = @@ -529,92 +514,34 @@ let close_section () = let full_olddir = fst !path_prefix in pop_path_prefix (); add_entry oname (ClosedSection (List.rev (mark::secdecls))); - if !Flags.xml_export then !xml_close_section (basename (fst oname)); let newdecls = List.map discharge_item secdecls in Summary.unfreeze_summaries fs; List.iter (Option.iter (fun (id,o) -> add_discharged_leaf id o)) newdecls; - Cooking.clear_cooking_sharing (); Nametab.push_dir (Nametab.Until 1) full_olddir (DirClosedSection full_olddir) -(*****************) -(* Backtracking. *) - -let (inLabel : int -> obj), (outLabel : obj -> int) = - declare_object_full {(default_object "DOT") with - classify_function = (fun _ -> Dispose)} - -let recache_decl = function - | (sp, Leaf o) -> cache_object (sp,o) - | (_,OpenedSection _) -> add_section () - | _ -> () - -let recache_context ctx = - List.iter recache_decl ctx - -let is_frozen_state = function (_,FrozenState _) -> true | _ -> false - -let set_lib_stk new_lib_stk = - lib_stk := new_lib_stk; - recalc_path_prefix (); - let spf = match find_entry_p is_frozen_state with - | (sp, FrozenState f) -> unfreeze_summaries f; sp - | _ -> assert false - in - let (after,_,_) = split_lib spf in - try - recache_context after - with - | Not_found -> error "Tried to set environment to an incoherent state." - -let reset_to_gen test = - let (_,_,before) = split_lib_gen test in - set_lib_stk before - -let reset_to sp = reset_to_gen (fun x -> fst x = sp) - -let first_command_label = 1 - -let mark_end_of_command, current_command_label, reset_command_label = - let n = ref (first_command_label-1) in - (fun () -> - match !lib_stk with - (_,Leaf o)::_ when object_tag o = "DOT" -> () - | _ -> incr n;add_anonymous_leaf (inLabel !n)), - (fun () -> !n), - (fun x -> n:=x;add_anonymous_leaf (inLabel x)) - -let is_label_n n x = - match x with - | (sp,Leaf o) when object_tag o = "DOT" && n = outLabel o -> true - | _ -> false - -(** Reset the label registered by [mark_end_of_command()] with number n, - which should be strictly in the past. *) - -let reset_label n = - if n >= current_command_label () then - error "Cannot backtrack to the current label or a future one"; - reset_to_gen (is_label_n n); - (* forget state numbers after n only if reset succeeded *) - reset_command_label n - -(** Search the last label registered before defining [id] *) - -let label_before_name (loc,id) = - let found = ref false in - let search = function - | (_,Leaf o) when !found && object_tag o = "DOT" -> true - | (sp,_) -> (if id = snd (repr_path (fst sp)) then found := true); false - in - match find_entry_p search with - | (_,Leaf o) -> outLabel o - | _ -> raise Not_found - (* State and initialization. *) -type frozen = Names.dir_path option * library_segment - -let freeze () = (!comp_name, !lib_stk) +type frozen = Names.DirPath.t option * library_segment + +let freeze ~marshallable = + match marshallable with + | `Shallow -> + (* TASSI: we should do something more sensible here *) + let lib_stk = + CList.map_filter (function + | _, Leaf _ -> None + | n, (CompilingLibrary _ as x) -> Some (n,x) + | n, OpenedModule (it,e,op,_) -> + Some(n,OpenedModule(it,e,op,Summary.empty_frozen)) + | n, ClosedModule _ -> Some (n,ClosedModule []) + | n, OpenedSection (op, _) -> + Some(n,OpenedSection(op,Summary.empty_frozen)) + | n, ClosedSection _ -> Some (n,ClosedSection []) + | _, FrozenState _ -> None) + !lib_stk in + !comp_name, lib_stk + | _ -> + !comp_name, !lib_stk let unfreeze (mn,stk) = comp_name := mn; @@ -622,57 +549,49 @@ let unfreeze (mn,stk) = recalc_path_prefix () let init () = - lib_stk := []; - add_frozen_state (); - comp_name := None; - path_prefix := initial_prefix; - init_summaries() + unfreeze (None,[]); + Summary.init_summaries (); + add_frozen_state () (* Stores e.g. the keywords declared in g_*.ml4 *) (* Misc *) -let mp_of_global ref = - match ref with - | VarRef id -> fst (current_prefix ()) - | ConstRef cst -> Names.con_modpath cst - | IndRef ind -> Names.ind_modpath ind - | ConstructRef constr -> Names.constr_modpath constr - -let rec dp_of_mp modp = - match modp with - | Names.MPfile dp -> dp - | Names.MPbound _ -> library_dp () - | Names.MPdot (mp,_) -> dp_of_mp mp - -let rec split_mp mp = - match mp with - | Names.MPfile dp -> dp, Names.empty_dirpath - | Names.MPdot (prfx, lbl) -> - let mprec, dprec = split_mp prfx in - mprec, Names.make_dirpath (Names.id_of_string (Names.string_of_label lbl) :: (Names.repr_dirpath dprec)) - | Names.MPbound mbid -> let (_, id, dp) = Names.repr_mbid mbid in library_dp(), Names.make_dirpath [id] - -let split_modpath mp = - let rec aux = function - | Names.MPfile dp -> dp, [] - | Names.MPbound mbid -> - library_dp (), [Names.id_of_mbid mbid] - | Names.MPdot (mp,l) -> let (mp', lab) = aux mp in - (mp', Names.id_of_label l :: lab) - in - let (mp, l) = aux mp in - mp, l - -let library_part ref = - match ref with - | VarRef id -> library_dp () - | _ -> dp_of_mp (mp_of_global ref) +let mp_of_global = function + |VarRef id -> current_mp () + |ConstRef cst -> Names.con_modpath cst + |IndRef ind -> Names.ind_modpath ind + |ConstructRef constr -> Names.constr_modpath constr + +let rec dp_of_mp = function + |Names.MPfile dp -> dp + |Names.MPbound _ -> library_dp () + |Names.MPdot (mp,_) -> dp_of_mp mp + +let rec split_mp = function + |Names.MPfile dp -> dp, Names.DirPath.empty + |Names.MPdot (prfx, lbl) -> + let mprec, dprec = split_mp prfx in + mprec, Libnames.add_dirpath_suffix dprec (Names.Label.to_id lbl) + |Names.MPbound mbid -> + let (_,id,dp) = Names.MBId.repr mbid in + library_dp (), Names.DirPath.make [id] + +let rec split_modpath = function + |Names.MPfile dp -> dp, [] + |Names.MPbound mbid -> library_dp (), [Names.MBId.to_id mbid] + |Names.MPdot (mp,l) -> + let (dp,ids) = split_modpath mp in + (dp, Names.Label.to_id l :: ids) + +let library_part = function + |VarRef id -> library_dp () + |ref -> dp_of_mp (mp_of_global ref) let remove_section_part ref = let sp = Nametab.path_of_global ref in let dir,_ = repr_path sp in match ref with | VarRef id -> - anomaly "remove_section_part not supported on local variables" + anomaly (Pp.str "remove_section_part not supported on local variables") | _ -> if is_dirpath_prefix_of dir (cwd ()) then (* Not yet (fully) discharged *) @@ -684,36 +603,30 @@ let remove_section_part ref = (************************) (* Discharging names *) -let pop_kn kn = - let (mp,dir,l) = Names.repr_mind kn in - Names.make_mind mp (pop_dirpath dir) l - -let pop_con con = - let (mp,dir,l) = Names.repr_con con in - Names.make_con mp (pop_dirpath dir) l - let con_defined_in_sec kn = let _,dir,_ = Names.repr_con kn in - dir <> Names.empty_dirpath && fst (split_dirpath dir) = snd (current_prefix ()) + not (Names.DirPath.is_empty dir) && + Names.DirPath.equal (pop_dirpath dir) (current_sections ()) let defined_in_sec kn = let _,dir,_ = Names.repr_mind kn in - dir <> Names.empty_dirpath && fst (split_dirpath dir) = snd (current_prefix ()) + not (Names.DirPath.is_empty dir) && + Names.DirPath.equal (pop_dirpath dir) (current_sections ()) let discharge_global = function | ConstRef kn when con_defined_in_sec kn -> - ConstRef (pop_con kn) + ConstRef (Globnames.pop_con kn) | IndRef (kn,i) when defined_in_sec kn -> - IndRef (pop_kn kn,i) + IndRef (Globnames.pop_kn kn,i) | ConstructRef ((kn,i),j) when defined_in_sec kn -> - ConstructRef ((pop_kn kn,i),j) + ConstructRef ((Globnames.pop_kn kn,i),j) | r -> r let discharge_kn kn = - if defined_in_sec kn then pop_kn kn else kn + if defined_in_sec kn then Globnames.pop_kn kn else kn let discharge_con cst = - if con_defined_in_sec cst then pop_con cst else cst + if con_defined_in_sec cst then Globnames.pop_con cst else cst let discharge_inductive (kn,i) = (discharge_kn kn,i) diff --git a/library/lib.mli b/library/lib.mli index d546d1ff..9c4d26c5 100644 --- a/library/lib.mli +++ b/library/lib.mli @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* Libobject.obj -> Libnames.object_name +val add_leaf : Names.Id.t -> Libobject.obj -> Libnames.object_name val add_anonymous_leaf : Libobject.obj -> unit (** this operation adds all objects with the same name and calls [load_object] for each of them *) -val add_leaves : Names.identifier -> Libobject.obj list -> Libnames.object_name +val add_leaves : Names.Id.t -> Libobject.obj list -> Libnames.object_name val add_frozen_state : unit -> unit (** {6 ... } *) + +(** The function [contents] gives access to the current entire segment *) + +val contents : unit -> library_segment + (** The function [contents_after] returns the current library segment, - starting from a given section path. If not given, the entire segment - is returned. *) + starting from a given section path. *) -val contents_after : Libnames.object_name option -> library_segment +val contents_after : Libnames.object_name -> library_segment (** {6 Functions relative to current path } *) (** User-side names *) -val cwd : unit -> Names.dir_path -val cwd_except_section : unit -> Names.dir_path -val current_dirpath : bool -> Names.dir_path (* false = except sections *) -val make_path : Names.identifier -> Libnames.full_path -val make_path_except_section : Names.identifier -> Libnames.full_path -val path_of_include : unit -> Libnames.full_path +val cwd : unit -> Names.DirPath.t +val cwd_except_section : unit -> Names.DirPath.t +val current_dirpath : bool -> Names.DirPath.t (* false = except sections *) +val make_path : Names.Id.t -> Libnames.full_path +val make_path_except_section : Names.Id.t -> Libnames.full_path (** Kernel-side names *) -val current_prefix : unit -> Names.module_path * Names.dir_path -val make_kn : Names.identifier -> Names.kernel_name -val make_con : Names.identifier -> Names.constant +val current_mp : unit -> Names.module_path +val make_kn : Names.Id.t -> Names.kernel_name (** Are we inside an opened section *) val sections_are_opened : unit -> bool @@ -91,11 +93,14 @@ val sections_depth : unit -> int (** Are we inside an opened module type *) val is_module_or_modtype : unit -> bool val is_modtype : unit -> bool +(* [is_modtype_strict] checks not only if we are in a module type, but + if the latest module started is a module type. *) +val is_modtype_strict : unit -> bool val is_module : unit -> bool val current_mod_id : unit -> Names.module_ident (** Returns the opening node of a given name *) -val find_opening_node : Names.identifier -> node +val find_opening_node : Names.Id.t -> node (** {6 Modules and module types } *) @@ -121,88 +126,68 @@ val end_modtype : (** {6 Compilation units } *) -val start_compilation : Names.dir_path -> Names.module_path -> unit -val end_compilation : Names.dir_path -> Libnames.object_prefix * library_segment +val start_compilation : Names.DirPath.t -> Names.module_path -> unit +val end_compilation_checks : Names.DirPath.t -> Libnames.object_name +val end_compilation : + Libnames.object_name-> Libnames.object_prefix * library_segment -(** The function [library_dp] returns the [dir_path] of the current +(** The function [library_dp] returns the [DirPath.t] of the current compiling library (or [default_library]) *) -val library_dp : unit -> Names.dir_path +val library_dp : unit -> Names.DirPath.t (** Extract the library part of a name even if in a section *) -val dp_of_mp : Names.module_path -> Names.dir_path -val split_mp : Names.module_path -> Names.dir_path * Names.dir_path -val split_modpath : Names.module_path -> Names.dir_path * Names.identifier list -val library_part : Libnames.global_reference -> Names.dir_path -val remove_section_part : Libnames.global_reference -> Names.dir_path +val dp_of_mp : Names.module_path -> Names.DirPath.t +val split_mp : Names.module_path -> Names.DirPath.t * Names.DirPath.t +val split_modpath : Names.module_path -> Names.DirPath.t * Names.Id.t list +val library_part : Globnames.global_reference -> Names.DirPath.t +val remove_section_part : Globnames.global_reference -> Names.DirPath.t (** {6 Sections } *) -val open_section : Names.identifier -> unit +val open_section : Names.Id.t -> unit val close_section : unit -> unit -(** {6 Backtracking } *) - -(** NB: The next commands are low-level ones, do not use them directly - otherwise the command history stack in [Backtrack] will be out-of-sync. - Also note that [reset_initial] is now [reset_label first_command_label] *) - -(** Adds a "dummy" entry in lib_stk with a unique new label number. *) -val mark_end_of_command : unit -> unit - -(** Returns the current label number *) -val current_command_label : unit -> int - -(** The first label number *) -val first_command_label : int - -(** [reset_label n] resets [lib_stk] to the label n registered by - [mark_end_of_command()]. It forgets anything registered after - this label. The label should be strictly in the past. *) -val reset_label : int -> unit - -(** search the label registered immediately before adding some definition *) -val label_before_name : Names.identifier Util.located -> int - (** {6 We can get and set the state of the operations (used in [States]). } *) type frozen -val freeze : unit -> frozen +val freeze : marshallable:Summary.marshallable -> frozen val unfreeze : frozen -> unit val init : unit -> unit -(** XML output hooks *) -val set_xml_open_section : (Names.identifier -> unit) -> unit -val set_xml_close_section : (Names.identifier -> unit) -> unit - -type binding_kind = Explicit | Implicit - (** {6 Section management for discharge } *) -type variable_info = Names.identifier * binding_kind * Term.constr option * Term.types -type variable_context = variable_info list +type variable_info = Names.Id.t * Decl_kinds.binding_kind * + Term.constr option * Term.types +type variable_context = variable_info list +type abstr_info = variable_context * Univ.universe_level_subst * Univ.UContext.t -val instance_from_variable_context : variable_context -> Names.identifier array -val named_of_variable_context : variable_context -> Sign.named_context +val instance_from_variable_context : variable_context -> Names.Id.t array +val named_of_variable_context : variable_context -> Context.named_context -val section_segment_of_constant : Names.constant -> variable_context -val section_segment_of_mutual_inductive: Names.mutual_inductive -> variable_context +val section_segment_of_constant : Names.constant -> abstr_info +val section_segment_of_mutual_inductive: Names.mutual_inductive -> abstr_info -val section_instance : Libnames.global_reference -> Names.identifier array -val is_in_section : Libnames.global_reference -> bool +val section_instance : Globnames.global_reference -> Univ.universe_instance * Names.Id.t array +val is_in_section : Globnames.global_reference -> bool -val add_section_variable : Names.identifier -> binding_kind -> unit +val add_section_variable : Names.Id.t -> Decl_kinds.binding_kind -> Decl_kinds.polymorphic -> Univ.universe_context_set -> unit -val add_section_constant : Names.constant -> Sign.named_context -> unit -val add_section_kn : Names.mutual_inductive -> Sign.named_context -> unit -val replacement_context : unit -> - (Names.identifier array Names.Cmap.t * Names.identifier array Names.Mindmap.t) +val add_section_constant : bool (* is_projection *) -> + Names.constant -> Context.named_context -> unit +val add_section_kn : Names.mutual_inductive -> Context.named_context -> unit +val replacement_context : unit -> Opaqueproof.work_list (** {6 Discharge: decrease the section level if in the current section } *) val discharge_kn : Names.mutual_inductive -> Names.mutual_inductive val discharge_con : Names.constant -> Names.constant -val discharge_global : Libnames.global_reference -> Libnames.global_reference +val discharge_global : Globnames.global_reference -> Globnames.global_reference val discharge_inductive : Names.inductive -> Names.inductive +(* discharging a constant in one go *) +val full_replacement_context : unit -> Opaqueproof.work_list list +val full_section_segment_of_constant : + Names.constant -> (Context.named_context -> Context.named_context) list + diff --git a/library/libnames.ml b/library/libnames.ml index db97eae9..f2a9d041 100644 --- a/library/libnames.ml +++ b/library/libnames.ml @@ -1,214 +1,112 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* true | _ -> false -let isConstRef = function ConstRef _ -> true | _ -> false -let isIndRef = function IndRef _ -> true | _ -> false -let isConstructRef = function ConstructRef _ -> true | _ -> false - -let eq_gr gr1 gr2 = - match gr1,gr2 with - | ConstRef con1, ConstRef con2 -> eq_constant con1 con2 - | IndRef kn1,IndRef kn2 -> eq_ind kn1 kn2 - | ConstructRef kn1,ConstructRef kn2 -> eq_constructor kn1 kn2 - | _,_ -> gr1=gr2 - -let destVarRef = function VarRef ind -> ind | _ -> failwith "destVarRef" -let destConstRef = function ConstRef ind -> ind | _ -> failwith "destConstRef" -let destIndRef = function IndRef ind -> ind | _ -> failwith "destIndRef" -let destConstructRef = function ConstructRef ind -> ind | _ -> failwith "destConstructRef" - -let subst_constructor subst ((kn,i),j as ref) = - let kn' = subst_ind subst kn in - if kn==kn' then ref, mkConstruct ref - else ((kn',i),j), mkConstruct ((kn',i),j) - -let subst_global subst ref = match ref with - | VarRef var -> ref, mkVar var - | ConstRef kn -> - let kn',t = subst_con subst kn in - if kn==kn' then ref, mkConst kn else ConstRef kn', t - | IndRef (kn,i) -> - let kn' = subst_ind subst kn in - if kn==kn' then ref, mkInd (kn,i) else IndRef(kn',i), mkInd (kn',i) - | ConstructRef ((kn,i),j as c) -> - let c',t = subst_constructor subst c in - if c'==c then ref,t else ConstructRef c', t - -let canonical_gr = function - | ConstRef con -> ConstRef(constant_of_kn(canonical_con con)) - | IndRef (kn,i) -> IndRef(mind_of_kn(canonical_mind kn),i) - | ConstructRef ((kn,i),j )-> ConstructRef((mind_of_kn(canonical_mind kn),i),j) - | VarRef id -> VarRef id - -let global_of_constr c = match kind_of_term c with - | Const sp -> ConstRef sp - | Ind ind_sp -> IndRef ind_sp - | Construct cstr_cp -> ConstructRef cstr_cp - | Var id -> VarRef id - | _ -> raise Not_found - -let constr_of_global = function - | VarRef id -> mkVar id - | ConstRef sp -> mkConst sp - | ConstructRef sp -> mkConstruct sp - | IndRef sp -> mkInd sp - -let constr_of_reference = constr_of_global -let reference_of_constr = global_of_constr - -let global_ord_gen fc fmi x y = - let ind_ord (indx,ix) (indy,iy) = - let c = Pervasives.compare ix iy in - if c = 0 then kn_ord (fmi indx) (fmi indy) else c - in - match x, y with - | ConstRef cx, ConstRef cy -> kn_ord (fc cx) (fc cy) - | IndRef indx, IndRef indy -> ind_ord indx indy - | ConstructRef (indx,jx), ConstructRef (indy,jy) -> - let c = Pervasives.compare jx jy in - if c = 0 then ind_ord indx indy else c - | _, _ -> Pervasives.compare x y - -let global_ord_can = global_ord_gen canonical_con canonical_mind -let global_ord_user = global_ord_gen user_con user_mind - -(* By default, [global_reference] are ordered on their canonical part *) - -module RefOrdered = struct - type t = global_reference - let compare = global_ord_can -end - -module RefOrdered_env = struct - type t = global_reference - let compare = global_ord_user -end - -module Refset = Set.Make(RefOrdered) -module Refmap = Map.Make(RefOrdered) - -(* Extended global references *) - -type syndef_name = kernel_name - -type extended_global_reference = - | TrueGlobal of global_reference - | SynDef of syndef_name - -(* We order [extended_global_reference] via their user part - (cf. pretty printer) *) - -module ExtRefOrdered = struct - type t = extended_global_reference - let compare x y = - match x, y with - | TrueGlobal rx, TrueGlobal ry -> global_ord_user rx ry - | SynDef knx, SynDef kny -> kn_ord knx kny - | _, _ -> Pervasives.compare x y -end (**********************************************) -let pr_dirpath sl = (str (string_of_dirpath sl)) +let pr_dirpath sl = (str (DirPath.to_string sl)) (*s Operations on dirpaths *) -(* Pop the last n module idents *) -let pop_dirpath_n n dir = - make_dirpath (list_skipn n (repr_dirpath dir)) +let split_dirpath d = match DirPath.repr d with + | id :: l -> DirPath.make l, id + | _ -> failwith "split_dirpath" -let pop_dirpath p = match repr_dirpath p with - | [] -> anomaly "dirpath_prefix: empty dirpath" - | _::l -> make_dirpath l +let pop_dirpath p = match DirPath.repr p with + | _::l -> DirPath.make l + | [] -> failwith "pop_dirpath" + +(* Pop the last n module idents *) +let pop_dirpath_n n dir = DirPath.make (List.skipn n (DirPath.repr dir)) let is_dirpath_prefix_of d1 d2 = - list_prefix_of (List.rev (repr_dirpath d1)) (List.rev (repr_dirpath d2)) + List.prefix_of Id.equal + (List.rev (DirPath.repr d1)) (List.rev (DirPath.repr d2)) let chop_dirpath n d = - let d1,d2 = list_chop n (List.rev (repr_dirpath d)) in - make_dirpath (List.rev d1), make_dirpath (List.rev d2) + let d1,d2 = List.chop n (List.rev (DirPath.repr d)) in + DirPath.make (List.rev d1), DirPath.make (List.rev d2) let drop_dirpath_prefix d1 d2 = - let d = Util.list_drop_prefix (List.rev (repr_dirpath d1)) (List.rev (repr_dirpath d2)) in - make_dirpath (List.rev d) - -let append_dirpath d1 d2 = make_dirpath (repr_dirpath d2 @ repr_dirpath d1) + let d = + List.drop_prefix Id.equal + (List.rev (DirPath.repr d1)) (List.rev (DirPath.repr d2)) + in + DirPath.make (List.rev d) -(* To know how qualified a name should be to be understood in the current env*) -let add_dirpath_prefix id d = make_dirpath (repr_dirpath d @ [id]) +let append_dirpath d1 d2 = DirPath.make (DirPath.repr d2 @ DirPath.repr d1) -let split_dirpath d = - let l = repr_dirpath d in (make_dirpath (List.tl l), List.hd l) +let add_dirpath_prefix id d = DirPath.make (DirPath.repr d @ [id]) -let add_dirpath_suffix p id = make_dirpath (id :: repr_dirpath p) +let add_dirpath_suffix p id = DirPath.make (id :: DirPath.repr p) (* parsing *) let parse_dir s = let len = String.length s in let rec decoupe_dirs dirs n = - if n = len && n > 0 then error (s ^ " is an invalid path."); + if Int.equal n len && n > 0 then error (s ^ " is an invalid path."); if n >= len then dirs else let pos = try String.index_from s n '.' with Not_found -> len in - if pos = n then error (s ^ " is an invalid path."); + if Int.equal pos n then error (s ^ " is an invalid path."); let dir = String.sub s n (pos-n) in - decoupe_dirs ((id_of_string dir)::dirs) (pos+1) + decoupe_dirs ((Id.of_string dir)::dirs) (pos+1) in decoupe_dirs [] 0 let dirpath_of_string s = - make_dirpath (if s = "" then [] else parse_dir s) - -let string_of_dirpath = Names.string_of_dirpath + let path = match s with + | "" -> [] + | _ -> parse_dir s + in + DirPath.make path +let string_of_dirpath = Names.DirPath.to_string -module Dirset = Set.Make(struct type t = dir_path let compare = compare end) -module Dirmap = Map.Make(struct type t = dir_path let compare = compare end) +module Dirset = Set.Make(DirPath) +module Dirmap = Map.Make(DirPath) (*s Section paths are absolute names *) type full_path = { - dirpath : dir_path ; - basename : identifier } + dirpath : DirPath.t ; + basename : Id.t } + +let dirpath sp = sp.dirpath +let basename sp = sp.basename let make_path pa id = { dirpath = pa; basename = id } let repr_path { dirpath = pa; basename = id } = (pa,id) +let eq_full_path p1 p2 = + Id.equal p1.basename p2.basename && + DirPath.equal p1.dirpath p2.dirpath + (* parsing and printing of section paths *) let string_of_path sp = let (sl,id) = repr_path sp in - if repr_dirpath sl = [] then string_of_id id - else (string_of_dirpath sl) ^ "." ^ (string_of_id id) + match DirPath.repr sl with + | [] -> Id.to_string id + | _ -> (DirPath.to_string sl) ^ "." ^ (Id.to_string id) let sp_ord sp1 sp2 = let (p1,id1) = repr_path sp1 and (p2,id2) = repr_path sp2 in - let p_bit = compare p1 p2 in - if p_bit = 0 then id_ord id1 id2 else p_bit + let p_bit = DirPath.compare p1 p2 in + if Int.equal p_bit 0 then Id.compare id1 id2 else p_bit module SpOrdered = struct @@ -218,9 +116,6 @@ module SpOrdered = module Spmap = Map.Make(SpOrdered) -let dirpath sp = let (p,_) = repr_path sp in p -let basename sp = let (_,id) = repr_path sp in id - let path_of_string s = try let dir, id = split_dirpath (dirpath_of_string s) in @@ -232,34 +127,8 @@ let pr_path sp = str (string_of_path sp) let restrict_path n sp = let dir, s = repr_path sp in - let dir' = list_firstn n (repr_dirpath dir) in - make_path (make_dirpath dir') s - -let encode_mind dir id = make_mind (MPfile dir) empty_dirpath (label_of_id id) - -let encode_con dir id = make_con (MPfile dir) empty_dirpath (label_of_id id) - -let decode_mind kn = - let rec dir_of_mp = function - | MPfile dir -> repr_dirpath dir - | MPbound mbid -> - let _,_,dp = repr_mbid mbid in - let id = id_of_mbid mbid in - id::(repr_dirpath dp) - | MPdot(mp,l) -> (id_of_label l)::(dir_of_mp mp) - in - let mp,sec_dir,l = repr_mind kn in - if (repr_dirpath sec_dir) = [] then - (make_dirpath (dir_of_mp mp)),id_of_label l - else - anomaly "Section part should be empty!" - -let decode_con kn = - let mp,sec_dir,l = repr_con kn in - match mp,(repr_dirpath sec_dir) with - MPfile dir,[] -> (dir,id_of_label l) - | _ , [] -> anomaly "MPfile expected!" - | _ -> anomaly "Section part should be empty!" + let dir' = List.firstn n (DirPath.repr dir) in + make_path (DirPath.make dir') s (*s qualified names *) type qualid = full_path @@ -267,44 +136,51 @@ type qualid = full_path let make_qualid = make_path let repr_qualid = repr_path +let qualid_eq = eq_full_path + let string_of_qualid = string_of_path let pr_qualid = pr_path let qualid_of_string = path_of_string let qualid_of_path sp = sp -let qualid_of_ident id = make_qualid empty_dirpath id +let qualid_of_ident id = make_qualid DirPath.empty id let qualid_of_dirpath dir = let (l,a) = split_dirpath dir in make_qualid l a type object_name = full_path * kernel_name -type object_prefix = dir_path * (module_path * dir_path) +type object_prefix = DirPath.t * (module_path * DirPath.t) let make_oname (dirpath,(mp,dir)) id = - make_path dirpath id, make_kn mp dir (label_of_id id) + make_path dirpath id, make_kn mp dir (Label.of_id id) -(* to this type are mapped dir_path's in the nametab *) +(* to this type are mapped DirPath.t's in the nametab *) type global_dir_reference = | DirOpenModule of object_prefix | DirOpenModtype of object_prefix | DirOpenSection of object_prefix | DirModule of object_prefix - | DirClosedSection of dir_path + | DirClosedSection of DirPath.t (* this won't last long I hope! *) -(* | ModRef mp -> - let mp' = subst_modpath subst mp in if mp==mp' then ref else - ModRef mp' - | ModTypeRef kn -> - let kn' = subst_kernel_name subst kn in if kn==kn' then ref else - ModTypeRef kn' -*) +let eq_op (d1, (mp1, p1)) (d2, (mp2, p2)) = + DirPath.equal d1 d2 && + DirPath.equal p1 p2 && + mp_eq mp1 mp2 + +let eq_global_dir_reference r1 r2 = match r1, r2 with +| DirOpenModule op1, DirOpenModule op2 -> eq_op op1 op2 +| DirOpenModtype op1, DirOpenModtype op2 -> eq_op op1 op2 +| DirOpenSection op1, DirOpenSection op2 -> eq_op op1 op2 +| DirModule op1, DirModule op2 -> eq_op op1 op2 +| DirClosedSection dp1, DirClosedSection dp2 -> DirPath.equal dp1 dp2 +| _ -> false type reference = - | Qualid of qualid located - | Ident of identifier located + | Qualid of qualid Loc.located + | Ident of Id.t Loc.located let qualid_of_reference = function | Qualid (loc,qid) -> loc, qid @@ -312,31 +188,44 @@ let qualid_of_reference = function let string_of_reference = function | Qualid (loc,qid) -> string_of_qualid qid - | Ident (loc,id) -> string_of_id id + | Ident (loc,id) -> Id.to_string id let pr_reference = function | Qualid (_,qid) -> pr_qualid qid - | Ident (_,id) -> pr_id id + | Ident (_,id) -> str (Id.to_string id) let loc_of_reference = function | Qualid (loc,qid) -> loc | Ident (loc,id) -> loc -(* popping one level of section in global names *) - -let pop_con con = - let (mp,dir,l) = repr_con con in - Names.make_con mp (pop_dirpath dir) l - -let pop_kn kn = - let (mp,dir,l) = repr_mind kn in - Names.make_mind mp (pop_dirpath dir) l - -let pop_global_reference = function - | ConstRef con -> ConstRef (pop_con con) - | IndRef (kn,i) -> IndRef (pop_kn kn,i) - | ConstructRef ((kn,i),j) -> ConstructRef ((pop_kn kn,i),j) - | VarRef id -> anomaly "VarRef not poppable" +let eq_reference r1 r2 = match r1, r2 with +| Qualid (_, q1), Qualid (_, q2) -> qualid_eq q1 q2 +| Ident (_, id1), Ident (_, id2) -> Id.equal id1 id2 +| _ -> false + +let join_reference ns r = + match ns , r with + Qualid (_, q1), Qualid (loc, q2) -> + let (dp1,id1) = repr_qualid q1 in + let (dp2,id2) = repr_qualid q2 in + Qualid (loc, + make_qualid + (append_dirpath (append_dirpath dp1 (dirpath_of_string (Names.Id.to_string id1))) dp2) + id2) + | Qualid (_, q1), Ident (loc, id2) -> + let (dp1,id1) = repr_qualid q1 in + Qualid (loc, + make_qualid + (append_dirpath dp1 (dirpath_of_string (Names.Id.to_string id1))) + id2) + | Ident (_, id1), Qualid (loc, q2) -> + let (dp2,id2) = repr_qualid q2 in + Qualid (loc, make_qualid + (append_dirpath (dirpath_of_string (Names.Id.to_string id1)) dp2) + id2) + | Ident (_, id1), Ident (loc, id2) -> + Qualid (loc, make_qualid + (dirpath_of_string (Names.Id.to_string id1)) id2) (* Deprecated synonyms *) diff --git a/library/libnames.mli b/library/libnames.mli index 8d026f42..3b5feb94 100644 --- a/library/libnames.mli +++ b/library/libnames.mli @@ -1,115 +1,57 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* bool -val isConstRef : global_reference -> bool -val isIndRef : global_reference -> bool -val isConstructRef : global_reference -> bool - -val eq_gr : global_reference -> global_reference -> bool -val canonical_gr : global_reference -> global_reference - -val destVarRef : global_reference -> variable -val destConstRef : global_reference -> constant -val destIndRef : global_reference -> inductive -val destConstructRef : global_reference -> constructor - - -val subst_constructor : substitution -> constructor -> constructor * constr -val subst_global : substitution -> global_reference -> global_reference * constr - -(** Turn a global reference into a construction *) -val constr_of_global : global_reference -> constr - -(** Turn a construction denoting a global reference into a global reference; - raise [Not_found] if not a global reference *) -val global_of_constr : constr -> global_reference - -(** Obsolete synonyms for constr_of_global and global_of_constr *) -val constr_of_reference : global_reference -> constr -val reference_of_constr : constr -> global_reference - -module RefOrdered : sig - type t = global_reference - val compare : global_reference -> global_reference -> int -end - -module RefOrdered_env : sig - type t = global_reference - val compare : global_reference -> global_reference -> int -end - -module Refset : Set.S with type elt = global_reference -module Refmap : Map.S with type key = global_reference - -(** {6 Extended global references } *) - -type syndef_name = kernel_name - -type extended_global_reference = - | TrueGlobal of global_reference - | SynDef of syndef_name - -module ExtRefOrdered : sig - type t = extended_global_reference - val compare : t -> t -> int -end (** {6 Dirpaths } *) -val pr_dirpath : dir_path -> Pp.std_ppcmds +(** FIXME: ought to be in Names.dir_path *) -val dirpath_of_string : string -> dir_path -val string_of_dirpath : dir_path -> string +val pr_dirpath : DirPath.t -> Pp.std_ppcmds -(** Pop the suffix of a [dir_path] *) -val pop_dirpath : dir_path -> dir_path +val dirpath_of_string : string -> DirPath.t +val string_of_dirpath : DirPath.t -> string + +(** Pop the suffix of a [DirPath.t]. Raises a [Failure] for an empty path *) +val pop_dirpath : DirPath.t -> DirPath.t (** Pop the suffix n times *) -val pop_dirpath_n : int -> dir_path -> dir_path +val pop_dirpath_n : int -> DirPath.t -> DirPath.t -(** Give the immediate prefix and basename of a [dir_path] *) -val split_dirpath : dir_path -> dir_path * identifier +(** Immediate prefix and basename of a [DirPath.t]. May raise [Failure] *) +val split_dirpath : DirPath.t -> DirPath.t * Id.t -val add_dirpath_suffix : dir_path -> module_ident -> dir_path -val add_dirpath_prefix : module_ident -> dir_path -> dir_path +val add_dirpath_suffix : DirPath.t -> module_ident -> DirPath.t +val add_dirpath_prefix : module_ident -> DirPath.t -> DirPath.t -val chop_dirpath : int -> dir_path -> dir_path * dir_path -val append_dirpath : dir_path -> dir_path -> dir_path +val chop_dirpath : int -> DirPath.t -> DirPath.t * DirPath.t +val append_dirpath : DirPath.t -> DirPath.t -> DirPath.t -val drop_dirpath_prefix : dir_path -> dir_path -> dir_path -val is_dirpath_prefix_of : dir_path -> dir_path -> bool +val drop_dirpath_prefix : DirPath.t -> DirPath.t -> DirPath.t +val is_dirpath_prefix_of : DirPath.t -> DirPath.t -> bool -module Dirset : Set.S with type elt = dir_path -module Dirmap : Map.S with type key = dir_path +module Dirset : Set.S with type elt = DirPath.t +module Dirmap : Map.ExtS with type key = DirPath.t and module Set := Dirset (** {6 Full paths are {e absolute} paths of declarations } *) type full_path +val eq_full_path : full_path -> full_path -> bool + (** Constructors of [full_path] *) -val make_path : dir_path -> identifier -> full_path +val make_path : DirPath.t -> Id.t -> full_path (** Destructors of [full_path] *) -val repr_path : full_path -> dir_path * identifier -val dirpath : full_path -> dir_path -val basename : full_path -> identifier +val repr_path : full_path -> DirPath.t * Id.t +val dirpath : full_path -> DirPath.t +val basename : full_path -> Id.t (** Parsing and printing of section path as ["coq_root.module.id"] *) val path_of_string : string -> full_path @@ -120,14 +62,6 @@ module Spmap : Map.S with type key = full_path val restrict_path : int -> full_path -> full_path -(** {6 Temporary function to brutally form kernel names from section paths } *) - -val encode_mind : dir_path -> identifier -> mutual_inductive -val decode_mind : mutual_inductive -> dir_path * identifier -val encode_con : dir_path -> identifier -> constant -val decode_con : constant -> dir_path * identifier - - (** {6 ... } *) (** A [qualid] is a partially qualified ident; it includes fully qualified names (= absolute names) and all intermediate partial @@ -136,19 +70,21 @@ val decode_con : constant -> dir_path * identifier type qualid -val make_qualid : dir_path -> identifier -> qualid -val repr_qualid : qualid -> dir_path * identifier +val make_qualid : DirPath.t -> Id.t -> qualid +val repr_qualid : qualid -> DirPath.t * Id.t + +val qualid_eq : qualid -> qualid -> bool val pr_qualid : qualid -> std_ppcmds val string_of_qualid : qualid -> string val qualid_of_string : string -> qualid -(** Turns an absolute name, a dirpath, or an identifier into a +(** Turns an absolute name, a dirpath, or an Id.t into a qualified name denoting the same name *) val qualid_of_path : full_path -> qualid -val qualid_of_dirpath : dir_path -> qualid -val qualid_of_ident : identifier -> qualid +val qualid_of_dirpath : DirPath.t -> qualid +val qualid_of_ident : Id.t -> qualid (** Both names are passed to objects: a "semantic" [kernel_name], which can be substituted and a "syntactic" [full_path] which can be printed @@ -156,19 +92,24 @@ val qualid_of_ident : identifier -> qualid type object_name = full_path * kernel_name -type object_prefix = dir_path * (module_path * dir_path) +type object_prefix = DirPath.t * (module_path * DirPath.t) + +val eq_op : object_prefix -> object_prefix -> bool -val make_oname : object_prefix -> identifier -> object_name +val make_oname : object_prefix -> Id.t -> object_name -(** to this type are mapped [dir_path]'s in the nametab *) +(** to this type are mapped [DirPath.t]'s in the nametab *) type global_dir_reference = | DirOpenModule of object_prefix | DirOpenModtype of object_prefix | DirOpenSection of object_prefix | DirModule of object_prefix - | DirClosedSection of dir_path + | DirClosedSection of DirPath.t (** this won't last long I hope! *) +val eq_global_dir_reference : + global_dir_reference -> global_dir_reference -> bool + (** {6 ... } *) (** A [reference] is the user-level notion of name. It denotes either a global name (referred either by a qualified name or by a single @@ -176,20 +117,16 @@ type global_dir_reference = type reference = | Qualid of qualid located - | Ident of identifier located + | Ident of Id.t located +val eq_reference : reference -> reference -> bool val qualid_of_reference : reference -> qualid located val string_of_reference : reference -> string val pr_reference : reference -> std_ppcmds -val loc_of_reference : reference -> loc - -(** {6 Popping one level of section in global names } *) - -val pop_con : constant -> constant -val pop_kn : mutual_inductive-> mutual_inductive -val pop_global_reference : global_reference -> global_reference +val loc_of_reference : reference -> Loc.t +val join_reference : reference -> reference -> reference (** Deprecated synonyms *) -val make_short_qualid : identifier -> qualid (** = qualid_of_ident *) +val make_short_qualid : Id.t -> qualid (** = qualid_of_ident *) val qualid_of_sp : full_path -> qualid (** = qualid_of_path *) diff --git a/library/libobject.ml b/library/libobject.ml index ee1c94b9..5f2a2127 100644 --- a/library/libobject.ml +++ b/library/libobject.ml @@ -1,15 +1,12 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* object_name * 'a -> unit; open_function : int -> object_name * 'a -> unit; classify_function : 'a -> 'a substitutivity; - subst_function : substitution * 'a -> 'a; + subst_function : Mod_subst.substitution * 'a -> 'a; discharge_function : object_name * 'a -> 'a option; rebuild_function : 'a -> 'a } -let yell s = anomaly s +let yell s = Errors.anomaly (Pp.str s) let default_object s = { object_name = s; @@ -69,12 +66,13 @@ type dynamic_object_declaration = { dyn_cache_function : object_name * obj -> unit; dyn_load_function : int -> object_name * obj -> unit; dyn_open_function : int -> object_name * obj -> unit; - dyn_subst_function : substitution * obj -> obj; + dyn_subst_function : Mod_subst.substitution * obj -> obj; dyn_classify_function : obj -> obj substitutivity; dyn_discharge_function : object_name * obj -> obj option; dyn_rebuild_function : obj -> obj } -let object_tag lobj = Dyn.tag lobj +let object_tag = Dyn.tag +let object_has_tag = Dyn.has_tag let cache_tab = (Hashtbl.create 17 : (string,dynamic_object_declaration) Hashtbl.t) @@ -82,36 +80,18 @@ let cache_tab = let declare_object_full odecl = let na = odecl.object_name in let (infun,outfun) = Dyn.create na in - let cacher (oname,lobj) = - if Dyn.tag lobj = na then odecl.cache_function (oname,outfun lobj) - else anomaly "somehow we got the wrong dynamic object in the cachefun" - and loader i (oname,lobj) = - if Dyn.tag lobj = na then odecl.load_function i (oname,outfun lobj) - else anomaly "somehow we got the wrong dynamic object in the loadfun" - and opener i (oname,lobj) = - if Dyn.tag lobj = na then odecl.open_function i (oname,outfun lobj) - else anomaly "somehow we got the wrong dynamic object in the openfun" - and substituter (sub,lobj) = - if Dyn.tag lobj = na then - infun (odecl.subst_function (sub,outfun lobj)) - else anomaly "somehow we got the wrong dynamic object in the substfun" - and classifier lobj = - if Dyn.tag lobj = na then - match odecl.classify_function (outfun lobj) with - | Dispose -> Dispose - | Substitute obj -> Substitute (infun obj) - | Keep obj -> Keep (infun obj) - | Anticipate (obj) -> Anticipate (infun obj) - else - anomaly "somehow we got the wrong dynamic object in the classifyfun" + let cacher (oname,lobj) = odecl.cache_function (oname,outfun lobj) + and loader i (oname,lobj) = odecl.load_function i (oname,outfun lobj) + and opener i (oname,lobj) = odecl.open_function i (oname,outfun lobj) + and substituter (sub,lobj) = infun (odecl.subst_function (sub,outfun lobj)) + and classifier lobj = match odecl.classify_function (outfun lobj) with + | Dispose -> Dispose + | Substitute obj -> Substitute (infun obj) + | Keep obj -> Keep (infun obj) + | Anticipate (obj) -> Anticipate (infun obj) and discharge (oname,lobj) = - if Dyn.tag lobj = na then - Option.map infun (odecl.discharge_function (oname,outfun lobj)) - else - anomaly "somehow we got the wrong dynamic object in the dischargefun" - and rebuild lobj = - if Dyn.tag lobj = na then infun (odecl.rebuild_function (outfun lobj)) - else anomaly "somehow we got the wrong dynamic object in the rebuildfun" + Option.map infun (odecl.discharge_function (oname,outfun lobj)) + and rebuild lobj = infun (odecl.rebuild_function (outfun lobj)) in Hashtbl.add cache_tab na { dyn_cache_function = cacher; dyn_load_function = loader; diff --git a/library/libobject.mli b/library/libobject.mli index 9e41b7fa..09938189 100644 --- a/library/libobject.mli +++ b/library/libobject.mli @@ -1,12 +1,11 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* ('a -> obj) val object_tag : obj -> string +val object_has_tag : obj -> string -> bool val cache_object : object_name * obj -> unit val load_object : int -> object_name * obj -> unit diff --git a/library/library.ml b/library/library.ml index 681c55c7..b078e2c4 100644 --- a/library/library.ml +++ b/library/library.ml @@ -1,182 +1,62 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* p = phys_dir) !load_paths with - | [_,dir,_] -> dir - | [] -> Nameops.default_root_prefix - | l -> anomaly ("Two logical paths are associated to "^phys_dir) - -let is_in_load_paths phys_dir = - let dir = System.canonical_path_name phys_dir in - let lp = get_load_paths () in - let check_p = fun p -> (String.compare dir p) == 0 in - List.exists check_p lp - -let remove_load_path dir = - load_paths := List.filter (fun (p,d,_) -> p <> dir) !load_paths - -let add_load_path isroot (phys_path,coq_path) = - let phys_path = System.canonical_path_name phys_path in - match List.filter (fun (p,d,_) -> p = phys_path) !load_paths with - | [_,dir,_] -> - if coq_path <> dir - (* If this is not the default -I . to coqtop *) - && not - (phys_path = System.canonical_path_name Filename.current_dir_name - && coq_path = Nameops.default_root_prefix) - then - begin - (* Assume the user is concerned by library naming *) - if dir <> Nameops.default_root_prefix then - Flags.if_warn msg_warning - (str phys_path ++ strbrk " was previously bound to " ++ - pr_dirpath dir ++ strbrk "; it is remapped to " ++ - pr_dirpath coq_path); - remove_load_path phys_path; - load_paths := (phys_path,coq_path,isroot) :: !load_paths; - end - | [] -> - load_paths := (phys_path,coq_path,isroot) :: !load_paths; - | _ -> anomaly ("Two logical paths are associated to "^phys_path) - -let extend_path_with_dirpath p dir = - List.fold_left Filename.concat p - (List.map string_of_id (List.rev (repr_dirpath dir))) - -let root_paths_matching_dir_path dir = - let rec aux = function - | [] -> [] - | (p,d,true) :: l when is_dirpath_prefix_of d dir -> - let suffix = drop_dirpath_prefix d dir in - extend_path_with_dirpath p suffix :: aux l - | _ :: l -> aux l in - aux !load_paths - -(* Root p is bound to A.B.C.D and we require file C.D.E.F *) -(* We may mean A.B.C.D.E.F, or A.B.C.D.C.D.E.F *) - -(* Root p is bound to A.B.C.C and we require file C.C.E.F *) -(* We may mean A.B.C.C.E.F, or A.B.C.C.C.E.F, or A.B.C.C.C.C.E.F *) - -let intersections d1 d2 = - let rec aux d1 = - if d1 = empty_dirpath then [d2] else - let rest = aux (snd (chop_dirpath 1 d1)) in - if is_dirpath_prefix_of d1 d2 then drop_dirpath_prefix d1 d2 :: rest - else rest in - aux d1 - -let loadpaths_matching_dir_path dir = - let rec aux = function - | [] -> [] - | (p,d,true) :: l -> - let inters = intersections d dir in - List.map (fun tl -> (extend_path_with_dirpath p tl,append_dirpath d tl)) - inters @ - aux l - | (p,d,_) :: l -> - (extend_path_with_dirpath p dir,append_dirpath d dir) :: aux l in - aux !load_paths - -let get_full_load_paths () = List.map (fun (a,b,c) -> (a,b)) !load_paths (************************************************************************) (*s Modules on disk contain the following informations (after the magic number, and before the digest). *) -type compilation_unit_name = dir_path +type compilation_unit_name = DirPath.t type library_disk = { md_name : compilation_unit_name; - md_compiled : LightenLibrary.lightened_compiled_library; + md_compiled : Safe_typing.compiled_library; md_objects : Declaremods.library_objects; - md_deps : (compilation_unit_name * Digest.t) list; - md_imports : compilation_unit_name list } + md_deps : (compilation_unit_name * Safe_typing.vodigest) array; + md_imports : compilation_unit_name array } (*s Modules loaded in memory contain the following informations. They are kept in the global table [libraries_table]. *) type library_t = { library_name : compilation_unit_name; - library_compiled : compiled_library; + library_compiled : Safe_typing.compiled_library; library_objects : Declaremods.library_objects; - library_deps : (compilation_unit_name * Digest.t) list; - library_imports : compilation_unit_name list; - library_digest : Digest.t } - -module LibraryOrdered = - struct - type t = dir_path - let compare d1 d2 = - Pervasives.compare - (List.rev (repr_dirpath d1)) (List.rev (repr_dirpath d2)) - end + library_deps : (compilation_unit_name * Safe_typing.vodigest) array; + library_imports : compilation_unit_name array; + library_digests : Safe_typing.vodigest; + library_extra_univs : Univ.universe_context_set; +} +module LibraryOrdered = DirPath module LibraryMap = Map.Make(LibraryOrdered) module LibraryFilenameMap = Map.Make(LibraryOrdered) (* This is a map from names to loaded libraries *) -let libraries_table = ref LibraryMap.empty +let libraries_table = Summary.ref LibraryMap.empty ~name:"LIBRARY" (* This is the map of loaded libraries filename *) (* (not synchronized so as not to be caught in the states on disk) *) let libraries_filename_table = ref LibraryFilenameMap.empty (* These are the _ordered_ sets of loaded, imported and exported libraries *) -let libraries_loaded_list = ref [] -let libraries_imports_list = ref [] -let libraries_exports_list = ref [] - -let freeze () = - !libraries_table, - !libraries_loaded_list, - !libraries_imports_list, - !libraries_exports_list - -let unfreeze (mt,mo,mi,me) = - libraries_table := mt; - libraries_loaded_list := mo; - libraries_imports_list := mi; - libraries_exports_list := me - -let init () = - libraries_table := LibraryMap.empty; - libraries_loaded_list := []; - libraries_imports_list := []; - libraries_exports_list := [] - -let _ = - Summary.declare_summary "MODULES" - { Summary.freeze_function = freeze; - Summary.unfreeze_function = unfreeze; - Summary.init_function = init } +let libraries_loaded_list = Summary.ref [] ~name:"LIBRARY-LOAD" +let libraries_imports_list = Summary.ref [] ~name:"LIBRARY-IMPORT" +let libraries_exports_list = Summary.ref [] ~name:"LIBRARY-EXPORT" (* various requests to the tables *) @@ -186,7 +66,7 @@ let find_library dir = let try_find_library dir = try find_library dir with Not_found -> - error ("Unknown library " ^ (string_of_dirpath dir)) + error ("Unknown library " ^ (DirPath.to_string dir)) let register_library_filename dir f = (* Not synchronized: overwrite the previous binding if one existed *) @@ -209,7 +89,7 @@ let library_is_loaded dir = with Not_found -> false let library_is_opened dir = - List.exists (fun m -> m.library_name = dir) !libraries_imports_list + List.exists (fun m -> DirPath.equal m.library_name dir) !libraries_imports_list let loaded_libraries () = List.map (fun m -> m.library_name) !libraries_loaded_list @@ -221,9 +101,17 @@ let opened_libraries () = be performed first, thus the libraries_loaded_list ... *) let register_loaded_library m = + let link m = + let dirname = Filename.dirname (library_full_filename m.library_name) in + let prefix = Nativecode.mod_uid_of_dirpath m.library_name ^ "." in + let f = prefix ^ "cmo" in + let f = Dynlink.adapt_filename f in + if not !Flags.no_native_compiler then + Nativelib.link_library ~prefix ~dirname ~basename:f + in let rec aux = function - | [] -> [m] - | m'::_ as l when m'.library_name = m.library_name -> l + | [] -> link m; [m] + | m'::_ as l when DirPath.equal m'.library_name m.library_name -> l | m'::l' -> m' :: aux l' in libraries_loaded_list := aux !libraries_loaded_list; libraries_table := LibraryMap.add m.library_name m !libraries_table @@ -237,7 +125,7 @@ let register_loaded_library m = let rec remember_last_of_each l m = match l with | [] -> [m] - | m'::l' when m'.library_name = m.library_name -> remember_last_of_each l' m + | m'::l' when DirPath.equal m'.library_name m.library_name -> remember_last_of_each l' m | m'::l' -> m' :: remember_last_of_each l' m let register_open_library export m = @@ -251,14 +139,14 @@ let register_open_library export m = (* [open_library export explicit m] opens library [m] if not already opened _or_ if explicitly asked to be (re)opened *) -let eq_lib_name m1 m2 = m1.library_name = m2.library_name +let eq_lib_name m1 m2 = DirPath.equal m1.library_name m2.library_name let open_library export explicit_libs m = if (* Only libraries indirectly to open are not reopen *) (* Libraries explicitly mentionned by the user are always reopen *) List.exists (eq_lib_name m) explicit_libs - or not (library_is_opened m.library_name) + || not (library_is_opened m.library_name) then begin register_open_library export m; Declaremods.really_import_module (MPfile m.library_name) @@ -275,7 +163,7 @@ let open_libraries export modl = List.fold_left (fun l m -> let subimport = - List.fold_left + Array.fold_left (fun l m -> remember_last_of_each l (try_find_library m)) l m.library_imports in remember_last_of_each subimport m) @@ -287,7 +175,7 @@ let open_libraries export modl = (* import and export - synchronous operations*) let open_import i (_,(dir,export)) = - if i=1 then + if Int.equal i 1 then (* even if the library is already imported, we re-import it *) (* if not (library_is_opened dir) then *) open_libraries export [try_find_library dir] @@ -300,7 +188,7 @@ let subst_import (_,o) = o let classify_import (_,export as obj) = if export then Substitute obj else Dispose -let in_import : dir_path * bool -> obj = +let in_import : DirPath.t * bool -> obj = declare_object {(default_object "IMPORT LIBRARY") with cache_function = cache_import; open_function = open_import; @@ -314,7 +202,7 @@ let in_import : dir_path * bool -> obj = (*s Loading from disk to cache (preparation phase) *) let (raw_extern_library, raw_intern_library) = - System.raw_extern_intern Coq_config.vo_magic_number ".vo" + System.raw_extern_intern Coq_config.vo_magic_number (************************************************************************) (*s Locate absolute or partially qualified library names in the path *) @@ -326,130 +214,233 @@ type library_location = LibLoaded | LibInPath let locate_absolute_library dir = (* Search in loadpath *) let pref, base = split_dirpath dir in - let loadpath = root_paths_matching_dir_path pref in - if loadpath = [] then raise LibUnmappedDir; - try - let name = (string_of_id base)^".vo" in - let _, file = System.where_in_path ~warn:false loadpath name in - (dir, file) - with Not_found -> - (* Last chance, removed from the file system but still in memory *) - if library_is_loaded dir then - (dir, library_full_filename dir) - else - raise LibNotFound + let loadpath = Loadpath.expand_root_path pref in + let () = match loadpath with [] -> raise LibUnmappedDir | _ -> () in + let find ext = + try + let name = Id.to_string base ^ ext in + let _, file = System.where_in_path ~warn:false loadpath name in + [file] + with Not_found -> [] in + match find ".vo" @ find ".vio" with + | [] -> raise LibNotFound + | [file] -> dir, file + | [vo;vi] when Unix.((stat vo).st_mtime < (stat vi).st_mtime) -> + msg_warning (str"Loading " ++ str vi ++ str " instead of " ++ + str vo ++ str " because it is more recent"); + dir, vi + | [vo;vi] -> dir, vo + | _ -> assert false let locate_qualified_library warn qid = - try - (* Search library in loadpath *) - let dir, base = repr_qualid qid in - let loadpath = loadpaths_matching_dir_path dir in - if loadpath = [] then raise LibUnmappedDir; - let name = string_of_id base ^ ".vo" in - let lpath, file = System.where_in_path ~warn (List.map fst loadpath) name in - let dir = add_dirpath_suffix (List.assoc lpath loadpath) base in - (* Look if loaded *) - if library_is_loaded dir then (LibLoaded, dir, library_full_filename dir) - (* Otherwise, look for it in the file system *) - else (LibInPath, dir, file) - with Not_found -> raise LibNotFound - -let explain_locate_library_error qid = function - | LibUnmappedDir -> - let prefix, _ = repr_qualid qid in - errorlabstrm "load_absolute_library_from" - (str "Cannot load " ++ pr_qualid qid ++ str ":" ++ spc () ++ - str "no physical path bound to" ++ spc () ++ pr_dirpath prefix ++ fnl ()) - | LibNotFound -> - errorlabstrm "load_absolute_library_from" - (str"Cannot find library " ++ pr_qualid qid ++ str" in loadpath") - | e -> raise e + (* Search library in loadpath *) + let dir, base = repr_qualid qid in + let loadpath = Loadpath.expand_path dir in + let () = match loadpath with [] -> raise LibUnmappedDir | _ -> () in + let find ext = + try + let name = Id.to_string base ^ ext in + let lpath, file = + System.where_in_path ~warn (List.map fst loadpath) name in + [lpath, file] + with Not_found -> [] in + let lpath, file = + match find ".vo" @ find ".vio" with + | [] -> raise LibNotFound + | [lpath, file] -> lpath, file + | [lpath_vo, vo; lpath_vi, vi] + when Unix.((stat vo).st_mtime < (stat vi).st_mtime) -> + msg_warning (str"Loading " ++ str vi ++ str " instead of " ++ + str vo ++ str " because it is more recent"); + lpath_vi, vi + | [lpath_vo, vo; _ ] -> lpath_vo, vo + | _ -> assert false + in + let dir = add_dirpath_suffix (String.List.assoc lpath loadpath) base in + (* Look if loaded *) + if library_is_loaded dir then (LibLoaded, dir,library_full_filename dir) + (* Otherwise, look for it in the file system *) + else (LibInPath, dir, file) + +let error_unmapped_dir qid = + let prefix, _ = repr_qualid qid in + errorlabstrm "load_absolute_library_from" + (str "Cannot load " ++ pr_qualid qid ++ str ":" ++ spc () ++ + str "no physical path bound to" ++ spc () ++ pr_dirpath prefix ++ fnl ()) + +let error_lib_not_found qid = + errorlabstrm "load_absolute_library_from" + (str"Cannot find library " ++ pr_qualid qid ++ str" in loadpath") let try_locate_absolute_library dir = try locate_absolute_library dir - with e when Errors.noncritical e -> - explain_locate_library_error (qualid_of_dirpath dir) e + with + | LibUnmappedDir -> error_unmapped_dir (qualid_of_dirpath dir) + | LibNotFound -> error_lib_not_found (qualid_of_dirpath dir) let try_locate_qualified_library (loc,qid) = try let (_,dir,f) = locate_qualified_library (Flags.is_verbose()) qid in dir,f + with + | LibUnmappedDir -> error_unmapped_dir qid + | LibNotFound -> error_lib_not_found qid + +(************************************************************************) +(** {6 Tables of opaque proof terms} *) + +(** We now store opaque proof terms apart from the rest of the environment. + See the [Indirect] contructor in [Lazyconstr.lazy_constr]. This way, + we can quickly load a first half of a .vo file without these opaque + terms, and access them only when a specific command (e.g. Print or + Print Assumptions) needs it. *) + +exception Faulty + +(** Fetching a table of opaque terms at position [pos] in file [f], + expecting to find first a copy of [digest]. *) + +let fetch_table what dp (f,pos,digest) = + let dir_path = Names.DirPath.to_string dp in + try + msg_info (str"Fetching " ++ str what++str" from disk for " ++ str dir_path); + let ch = System.with_magic_number_check raw_intern_library f in + let () = seek_in ch pos in + if not (String.equal (System.digest_in f ch) digest) then raise Faulty; + let table, pos', digest' = System.marshal_in_segment f ch in + let () = close_in ch in + let ch' = open_in f in + if not (String.equal (Digest.channel ch' pos') digest') then raise Faulty; + let () = close_in ch' in + table with e when Errors.noncritical e -> - explain_locate_library_error qid e + error + ("The file "^f^" (bound to " ^ dir_path ^ + ") is inaccessible or corrupted,\n" ^ + "cannot load some "^what^" in it.\n") + +(** Delayed / available tables of opaque terms *) + +type 'a table_status = + | ToFetch of string * int * Digest.t + | Fetched of 'a Future.computation array + +let opaque_tables = + ref (LibraryMap.empty : (Term.constr table_status) LibraryMap.t) +let univ_tables = + ref (LibraryMap.empty : (Univ.universe_context_set table_status) LibraryMap.t) + +let add_opaque_table dp st = + opaque_tables := LibraryMap.add dp st !opaque_tables +let add_univ_table dp st = + univ_tables := LibraryMap.add dp st !univ_tables + +let access_table fetch_table add_table tables dp i = + let t = match LibraryMap.find dp tables with + | Fetched t -> t + | ToFetch (f,pos,digest) -> + let t = fetch_table dp (f,pos,digest) in + add_table dp (Fetched t); + t + in + assert (i < Array.length t); t.(i) + +let access_opaque_table dp i = + access_table + (fetch_table "opaque proofs") + add_opaque_table !opaque_tables dp i +let access_univ_table dp i = + try + Some (access_table + (fetch_table "universe contexts of opaque proofs") + add_univ_table !univ_tables dp i) + with Not_found -> None +let () = + Opaqueproof.set_indirect_opaque_accessor access_opaque_table; + Opaqueproof.set_indirect_univ_accessor access_univ_table (************************************************************************) (* Internalise libraries *) -let mk_library md table digest = - let md_compiled = - LightenLibrary.load ~load_proof:!Flags.load_proofs table md.md_compiled - in { +type seg_lib = library_disk +type seg_univ = (* true = vivo, false = vi *) + Univ.universe_context_set Future.computation array * Univ.universe_context_set * bool +type seg_discharge = Opaqueproof.cooking_info list array +type seg_proofs = Term.constr Future.computation array + +let mk_library md digests univs = + { library_name = md.md_name; - library_compiled = md_compiled; + library_compiled = md.md_compiled; library_objects = md.md_objects; library_deps = md.md_deps; library_imports = md.md_imports; - library_digest = digest + library_digests = digests; + library_extra_univs = univs; } -let fetch_opaque_table (f,pos,digest) = - try - let ch = System.with_magic_number_check raw_intern_library f in - seek_in ch pos; - if System.marshal_in f ch <> digest then failwith "File changed!"; - let table = (System.marshal_in f ch : LightenLibrary.table) in - close_in ch; - table - with e when Errors.noncritical e -> - error - ("The file "^f^" is inaccessible or has changed,\n" ^ - "cannot load some opaque constant bodies in it.\n") - let intern_from_file f = let ch = System.with_magic_number_check raw_intern_library f in - let lmd = System.marshal_in f ch in - let pos = pos_in ch in - let digest = System.marshal_in f ch in - let table = lazy (fetch_opaque_table (f,pos,digest)) in - register_library_filename lmd.md_name f; - let library = mk_library lmd table digest in + let (lmd : seg_lib), pos, digest_lmd = System.marshal_in_segment f ch in + let (univs : seg_univ option), _, digest_u = System.marshal_in_segment f ch in + let _ = System.skip_in_segment f ch in + let pos, digest = System.skip_in_segment f ch in close_in ch; - library - -let rec intern_library needed (dir, f) = + register_library_filename lmd.md_name f; + add_opaque_table lmd.md_name (ToFetch (f,pos,digest)); + let open Safe_typing in + match univs with + | None -> mk_library lmd (Dvo_or_vi digest_lmd) Univ.ContextSet.empty + | Some (utab,uall,true) -> + add_univ_table lmd.md_name (Fetched utab); + mk_library lmd (Dvivo (digest_lmd,digest_u)) uall + | Some (utab,_,false) -> + add_univ_table lmd.md_name (Fetched utab); + mk_library lmd (Dvo_or_vi digest_lmd) Univ.ContextSet.empty + +module DPMap = Map.Make(DirPath) + +let deps_to_string deps = + Array.fold_left (fun s (n, _) -> s^"\n - "^(DirPath.to_string n)) "" deps + +let rec intern_library (needed, contents) (dir, f) from = + Pp.feedback(Feedback.FileDependency (from, f)); (* Look if in the current logical environment *) - try find_library dir, needed + try find_library dir, (needed, contents) with Not_found -> (* Look if already listed and consequently its dependencies too *) - try List.assoc dir needed, needed + try DPMap.find dir contents, (needed, contents) with Not_found -> (* [dir] is an absolute name which matches [f] which must be in loadpath *) let m = intern_from_file f in - if dir <> m.library_name then + if not (DirPath.equal dir m.library_name) then errorlabstrm "load_physical_library" (str ("The file " ^ f ^ " contains library") ++ spc () ++ pr_dirpath m.library_name ++ spc () ++ str "and not library" ++ spc() ++ pr_dirpath dir); - m, intern_library_deps needed dir m + Pp.feedback(Feedback.FileLoaded(DirPath.to_string dir, f)); + m, intern_library_deps (needed, contents) dir m (Some f) -and intern_library_deps needed dir m = - (dir,m)::List.fold_left (intern_mandatory_library dir) needed m.library_deps +and intern_library_deps libs dir m from = + let needed, contents = Array.fold_left (intern_mandatory_library dir from) libs m.library_deps in + (dir :: needed, DPMap.add dir m contents ) -and intern_mandatory_library caller needed (dir,d) = - let m,needed = intern_library needed (try_locate_absolute_library dir) in - if d <> m.library_digest then - errorlabstrm "" (strbrk ("Compiled library "^(string_of_dirpath caller)^ - ".vo makes inconsistent assumptions over library " - ^(string_of_dirpath dir))); - needed +and intern_mandatory_library caller from libs (dir,d) = + let m, libs = intern_library libs (try_locate_absolute_library dir) from in + if not (Safe_typing.digest_match ~actual:m.library_digests ~required:d) then + errorlabstrm "" (strbrk ("Compiled library "^ DirPath.to_string caller ^ + ".vo makes inconsistent assumptions over library " ^ + DirPath.to_string dir)); + libs -let rec_intern_library needed mref = - let _,needed = intern_library needed mref in needed +let rec_intern_library libs mref = + let _, libs = intern_library libs mref None in + libs let check_library_short_name f dir = function - | Some id when id <> snd (split_dirpath dir) -> + | Some id when not (Id.equal id (snd (split_dirpath dir))) -> errorlabstrm "check_library_short_name" (str ("The file " ^ f ^ " contains library") ++ spc () ++ pr_dirpath dir ++ spc () ++ str "and not library" ++ spc () ++ @@ -463,18 +454,24 @@ let rec_intern_by_filename_only id f = (* We check no other file containing same library is loaded *) if library_is_loaded m.library_name then begin - Flags.if_warn msg_warning + msg_warning (pr_dirpath m.library_name ++ str " is already loaded from file " ++ str (library_full_filename m.library_name)); m.library_name, [] end else - let needed = intern_library_deps [] m.library_name m in + let needed, contents = intern_library_deps ([], DPMap.empty) m.library_name m (Some f) in + let needed = List.map (fun dir -> dir, DPMap.find dir contents) needed in m.library_name, needed +let native_name_from_filename f = + let ch = System.with_magic_number_check raw_intern_library f in + let (lmd : seg_lib), pos, digest_lmd = System.marshal_in_segment f ch in + Nativecode.mod_uid_of_dirpath lmd.md_name + let rec_intern_library_from_file idopt f = (* A name is specified, we have to check it contains library id *) - let paths = get_load_paths () in + let paths = Loadpath.get_paths () in let _, f = System.find_file_in_path ~warn:(Flags.is_verbose()) paths (f^".vo") in rec_intern_by_filename_only idopt f @@ -496,14 +493,13 @@ let rec_intern_library_from_file idopt f = which recursively loads its dependencies) *) -type library_reference = dir_path list * bool option - let register_library m = Declaremods.register_library m.library_name m.library_compiled m.library_objects - m.library_digest; + m.library_digests + m.library_extra_univs; register_loaded_library m (* Follow the semantics of Anticipate object: @@ -526,7 +522,7 @@ let discharge_require (_,o) = Some o (* open_function is never called from here because an Anticipate object *) -type require_obj = library_t list * dir_path list * bool option +type require_obj = library_t list * DirPath.t list * bool option let in_require : require_obj -> obj = declare_object {(default_object "REQUIRE") with @@ -539,12 +535,9 @@ let in_require : require_obj -> obj = (* Require libraries, import them if [export <> None], mark them for export if [export = Some true] *) -let xml_require = ref (fun d -> ()) -let set_xml_require f = xml_require := f - let require_library_from_dirpath modrefl export = - let needed = List.fold_left rec_intern_library [] modrefl in - let needed = List.rev_map snd needed in + let needed, contents = List.fold_left rec_intern_library ([], DPMap.empty) modrefl in + let needed = List.rev_map (fun dir -> DPMap.find dir contents) needed in let modrefl = List.map fst modrefl in if Lib.is_module_or_modtype () then begin @@ -555,7 +548,6 @@ let require_library_from_dirpath modrefl export = end else add_anonymous_leaf (in_require (needed,modrefl,export)); - if !Flags.xml_export then List.iter !xml_require modrefl; add_frozen_state () let require_library qidl export = @@ -572,7 +564,6 @@ let require_library_from_file idopt file export = end else add_anonymous_leaf (in_require (needed,[modref],export)); - if !Flags.xml_export then !xml_require modref; add_frozen_state () (* the function called by Vernacentries.vernac_import *) @@ -597,28 +588,73 @@ let import_module export (loc,qid) = (*s Initializing the compilation of a library. *) let check_coq_overwriting p id = - let l = repr_dirpath p in - if not !Flags.boot && l <> [] && string_of_id (list_last l) = "Coq" then + let l = DirPath.repr p in + let is_empty = match l with [] -> true | _ -> false in + if not !Flags.boot && not is_empty && String.equal (Id.to_string (List.last l)) "Coq" then errorlabstrm "" - (strbrk ("Cannot build module "^string_of_dirpath p^"."^string_of_id id^ + (strbrk ("Cannot build module "^DirPath.to_string p^"."^Id.to_string id^ ": it starts with prefix \"Coq\" which is reserved for the Coq library.")) +(* Verifies that a string starts by a letter and do not contain + others caracters than letters, digits, or `_` *) + +let check_module_name s = + let msg c = + strbrk "Invalid module name: " ++ str s ++ strbrk " character " ++ + (if c = '\'' then str "\"'\"" else (str "'" ++ str (String.make 1 c) ++ str "'")) ++ + strbrk " is not allowed in module names\n" + in + let err c = errorlabstrm "" (msg c) in + match String.get s 0 with + | 'a' .. 'z' | 'A' .. 'Z' -> + for i = 1 to (String.length s)-1 do + match String.get s i with + | 'a' .. 'z' | 'A' .. 'Z' | '0' .. '9' | '_' -> () + | c -> err c + done + | c -> err c + let start_library f = - let paths = get_load_paths () in - let _,longf = + let paths = Loadpath.get_paths () in + let _, longf = System.find_file_in_path ~warn:(Flags.is_verbose()) paths (f^".v") in - let ldir0 = find_logical_path (Filename.dirname longf) in - let id = id_of_string (Filename.basename f) in + let ldir0 = + try + let lp = Loadpath.find_load_path (Filename.dirname longf) in + Loadpath.logical lp + with Not_found -> Nameops.default_root_prefix + in + let file = Filename.basename f in + let id = Id.of_string file in + check_module_name file; check_coq_overwriting ldir0 id; let ldir = add_dirpath_suffix ldir0 id in Declaremods.start_library ldir; ldir,longf +let load_library_todo f = + let paths = Loadpath.get_paths () in + let _, longf = + System.find_file_in_path ~warn:(Flags.is_verbose()) paths (f^".v") in + let f = longf^"io" in + let ch = System.with_magic_number_check raw_intern_library f in + let (s1 : seg_lib), _, _ = System.marshal_in_segment f ch in + let (s2 : seg_univ option), _, _ = System.marshal_in_segment f ch in + let (s3 : seg_discharge option), _, _ = System.marshal_in_segment f ch in + let tasks, _, _ = System.marshal_in_segment f ch in + let (s5 : seg_proofs), _, _ = System.marshal_in_segment f ch in + close_in ch; + if tasks = None then errorlabstrm "restart" (str"not a .vio file"); + if s2 = None then errorlabstrm "restart" (str"not a .vio file"); + if s3 = None then errorlabstrm "restart" (str"not a .vio file"); + if pi3 (Option.get s2) then errorlabstrm "restart" (str"not a .vio file"); + longf, s1, Option.get s2, Option.get s3, Option.get tasks, s5 + (************************************************************************) (*s [save_library dir] ends library [dir] and save it to the disk. *) let current_deps () = - List.map (fun m -> (m.library_name, m.library_digest)) !libraries_loaded_list + List.map (fun m -> m.library_name, m.library_digests) !libraries_loaded_list let current_reexports () = List.map (fun m -> m.library_name) !libraries_exports_list @@ -629,34 +665,85 @@ let error_recursively_dependent_library dir = strbrk " to save current library because" ++ strbrk " it already depends on a library of this name.") +(* We now use two different digests in a .vo file. The first one + only covers half of the file, without the opaque table. It is + used for identifying this version of this library : this digest + is the one leading to "inconsistent assumptions" messages. + The other digest comes at the very end, and covers everything + before it. This one is used for integrity check of the whole + file when loading the opaque table. *) + (* Security weakness: file might have been changed on disk between writing the content and computing the checksum... *) -let save_library_to dir f = - let cenv, seg = Declaremods.end_library dir in - let cenv, table = LightenLibrary.save cenv in + +let save_library_to ?todo dir f otab = + let f, except = match todo with + | None -> + assert(!Flags.compilation_mode = Flags.BuildVo); + f ^ "o", Future.UUIDSet.empty + | Some (l,_) -> + f ^ "io", + List.fold_left (fun e r -> Future.UUIDSet.add r.Stateid.uuid e) + Future.UUIDSet.empty l in + let cenv, seg, ast = Declaremods.end_library ~except dir in + let opaque_table, univ_table, disch_table, f2t_map = Opaqueproof.dump otab in + let tasks, utab, dtab = + match todo with + | None -> None, None, None + | Some (tasks, rcbackup) -> + let tasks = + List.map Stateid.(fun r -> + { r with uuid = Future.UUIDMap.find r.uuid f2t_map }) tasks in + Some (tasks,rcbackup), + Some (univ_table,Univ.ContextSet.empty,false), + Some disch_table in + let except = + Future.UUIDSet.fold (fun uuid acc -> + Int.Set.add (Future.UUIDMap.find uuid f2t_map) acc) + except Int.Set.empty in + let is_done_or_todo i x = Future.is_val x || Int.Set.mem i except in + Array.iteri (fun i x -> + if not(is_done_or_todo i x) then Errors.errorlabstrm "library" + Pp.(str"Proof object "++int i++str" is not checked nor to be checked")) + opaque_table; let md = { md_name = dir; md_compiled = cenv; md_objects = seg; - md_deps = current_deps (); - md_imports = current_reexports () } in - if List.mem_assoc dir md.md_deps then + md_deps = Array.of_list (current_deps ()); + md_imports = Array.of_list (current_reexports ()) } in + if Array.exists (fun (d,_) -> DirPath.equal d dir) md.md_deps then error_recursively_dependent_library dir; + (* Open the vo file and write the magic number *) let (f',ch) = raw_extern_library f in try - System.marshal_out ch md; - flush ch; - (* The loading of the opaque definitions table is optional whereas - the digest is loaded all the time. As a consequence, the digest - must be serialized before the table (if we want to keep the - current simple layout of .vo files). This also entails that the - digest does not take opaque terms into account anymore. *) - let di = Digest.file f' in - System.marshal_out ch di; - System.marshal_out ch table; - close_out ch - with reraise -> - msg_warn ("Removed file "^f'); close_out ch; Sys.remove f'; raise reraise + (* Writing vo payload *) + System.marshal_out_segment f' ch (md : seg_lib); + System.marshal_out_segment f' ch (utab : seg_univ option); + System.marshal_out_segment f' ch (dtab : seg_discharge option); + System.marshal_out_segment f' ch (tasks : 'tasks option); + System.marshal_out_segment f' ch (opaque_table : seg_proofs); + close_out ch; + (* Writing native code files *) + if not !Flags.no_native_compiler then + let fn = Filename.dirname f'^"/"^Nativecode.mod_uid_of_dirpath dir in + if not (Nativelib.compile_library dir ast fn) then + msg_error (str"Could not compile the library to native code. Skipping.") + with reraise -> + let reraise = Errors.push reraise in + let () = msg_warning (str ("Removed file "^f')) in + let () = close_out ch in + let () = Sys.remove f' in + iraise reraise + +let save_library_raw f lib univs proofs = + let (f',ch) = raw_extern_library (f^"o") in + System.marshal_out_segment f' ch (lib : seg_lib); + System.marshal_out_segment f' ch (Some univs : seg_univ option); + System.marshal_out_segment f' ch (None : seg_discharge option); + System.marshal_out_segment f' ch (None : 'tasks option); + System.marshal_out_segment f' ch (proofs : seg_proofs); + close_out ch (************************************************************************) (*s Display the memory use of a library. *) @@ -666,5 +753,16 @@ open Printf let mem s = let m = try_find_library s in h 0 (str (sprintf "%dk (cenv = %dk / seg = %dk)" - (size_kb m) (size_kb m.library_compiled) - (size_kb m.library_objects))) + (CObj.size_kb m) (CObj.size_kb m.library_compiled) + (CObj.size_kb m.library_objects))) + +module StringOrd = struct type t = string let compare = String.compare end +module StringSet = Set.Make(StringOrd) + +let get_used_load_paths () = + StringSet.elements + (List.fold_left (fun acc m -> StringSet.add + (Filename.dirname (library_full_filename m.library_name)) acc) + StringSet.empty !libraries_loaded_list) + +let _ = Nativelib.get_load_paths := get_used_load_paths diff --git a/library/library.mli b/library/library.mli index 630b9f58..13d83a5c 100644 --- a/library/library.mli +++ b/library/library.mli @@ -1,15 +1,14 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* bool option -> unit -val require_library_from_dirpath : (dir_path * string) list -> bool option -> unit +val require_library_from_dirpath : (DirPath.t * string) list -> bool option -> unit val require_library_from_file : - identifier option -> System.physical_path -> bool option -> unit + Id.t option -> CUnix.physical_path -> bool option -> unit (** {6 ... } *) + +(** Segments of a library *) +type seg_lib +type seg_univ = (* cst, all_cst, finished? *) + Univ.universe_context_set Future.computation array * Univ.universe_context_set * bool +type seg_discharge = Opaqueproof.cooking_info list array +type seg_proofs = Term.constr Future.computation array + (** Open a module (or a library); if the boolean is true then it's also an export otherwise just a simple import *) val import_module : bool -> qualid located -> unit (** {6 Start the compilation of a library } *) -val start_library : string -> dir_path * string +val start_library : string -> DirPath.t * string (** {6 End the compilation of a library and save it to a ".vo" file } *) -val save_library_to : dir_path -> string -> unit +val save_library_to : + ?todo:((Future.UUID.t,'document) Stateid.request list * 'counters) -> + DirPath.t -> string -> Opaqueproof.opaquetab -> unit + +val load_library_todo : + string -> string * seg_lib * seg_univ * seg_discharge * 'tasks * seg_proofs +val save_library_raw : string -> seg_lib -> seg_univ -> seg_proofs -> unit (** {6 Interrogate the status of libraries } *) (** - Tell if a library is loaded or opened *) -val library_is_loaded : dir_path -> bool -val library_is_opened : dir_path -> bool +val library_is_loaded : DirPath.t -> bool +val library_is_opened : DirPath.t -> bool (** - Tell which libraries are loaded or imported *) -val loaded_libraries : unit -> dir_path list -val opened_libraries : unit -> dir_path list +val loaded_libraries : unit -> DirPath.t list +val opened_libraries : unit -> DirPath.t list (** - Return the full filename of a loaded library. *) -val library_full_filename : dir_path -> string +val library_full_filename : DirPath.t -> string (** - Overwrite the filename of all libraries (used when restoring a state) *) val overwrite_library_filenames : string -> unit -(** {6 Hook for the xml exportation of libraries } *) -val set_xml_require : (dir_path -> unit) -> unit - -(** {6 ... } *) -(** Global load paths: a load path is a physical path in the file - system; to each load path is associated a Coq [dir_path] (the "logical" - path of the physical path) *) - -val get_load_paths : unit -> System.physical_path list -val get_full_load_paths : unit -> (System.physical_path * dir_path) list -val add_load_path : bool -> System.physical_path * dir_path -> unit -val remove_load_path : System.physical_path -> unit -val find_logical_path : System.physical_path -> dir_path -val is_in_load_paths : System.physical_path -> bool - (** {6 Locate a library in the load paths } *) exception LibUnmappedDir exception LibNotFound type library_location = LibLoaded | LibInPath val locate_qualified_library : - bool -> qualid -> library_location * dir_path * System.physical_path -val try_locate_qualified_library : qualid located -> dir_path * string + bool -> qualid -> library_location * DirPath.t * CUnix.physical_path +val try_locate_qualified_library : qualid located -> DirPath.t * string (** {6 Statistics: display the memory use of a library. } *) -val mem : dir_path -> Pp.std_ppcmds +val mem : DirPath.t -> Pp.std_ppcmds + +(** {6 Native compiler. } *) +val native_name_from_filename : string -> string diff --git a/library/library.mllib b/library/library.mllib index e8b5a7a4..eca28c82 100644 --- a/library/library.mllib +++ b/library/library.mllib @@ -1,16 +1,20 @@ Nameops Libnames +Globnames Libobject Summary Nametab Global +Universes Lib Declaremods +Loadpath Library States -Decl_kinds +Kindops Dischargedhypsmap Goptions Decls Heads -Assumptions \ No newline at end of file +Assumptions +Keys diff --git a/library/loadpath.ml b/library/loadpath.ml new file mode 100644 index 00000000..5876eedd --- /dev/null +++ b/library/loadpath.ml @@ -0,0 +1,135 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* raise Not_found + | [p] -> p + | _ -> anomaly_too_many_paths phys_dir + +let is_in_load_paths phys_dir = + let dir = CUnix.canonical_path_name phys_dir in + let lp = get_load_paths () in + let check_p p = String.equal dir p.path_physical in + List.exists check_p lp + +let remove_load_path dir = + let filter p = not (String.equal p.path_physical dir) in + load_paths := List.filter filter !load_paths + +let add_load_path phys_path path_type coq_path = + let phys_path = CUnix.canonical_path_name phys_path in + let filter p = String.equal p.path_physical phys_path in + let binding = { + path_logical = coq_path; + path_physical = phys_path; + path_type = path_type; + } in + match List.filter filter !load_paths with + | [] -> + load_paths := binding :: !load_paths + | [p] -> + let dir = p.path_logical in + if not (DirPath.equal coq_path dir) + (* If this is not the default -I . to coqtop *) + && not + (String.equal phys_path (CUnix.canonical_path_name Filename.current_dir_name) + && DirPath.equal coq_path (Nameops.default_root_prefix)) + then + begin + (* Assume the user is concerned by library naming *) + if not (DirPath.equal dir Nameops.default_root_prefix) then + msg_warning + (str phys_path ++ strbrk " was previously bound to " ++ + pr_dirpath dir ++ strbrk "; it is remapped to " ++ + pr_dirpath coq_path); + remove_load_path phys_path; + load_paths := binding :: !load_paths; + end + | _ -> anomaly_too_many_paths phys_path + +let extend_path_with_dirpath p dir = + List.fold_left Filename.concat p + (List.rev_map Id.to_string (DirPath.repr dir)) + +let expand_root_path dir = + let rec aux = function + | [] -> [] + | p :: l -> + if p.path_type <> ImplicitPath && is_dirpath_prefix_of p.path_logical dir then + let suffix = drop_dirpath_prefix p.path_logical dir in + extend_path_with_dirpath p.path_physical suffix :: aux l + else aux l + in + aux !load_paths + +(* Root p is bound to A.B.C.D and we require file C.D.E.F *) +(* We may mean A.B.C.D.E.F, or A.B.C.D.C.D.E.F *) + +(* Root p is bound to A.B.C.C and we require file C.C.E.F *) +(* We may mean A.B.C.C.E.F, or A.B.C.C.C.E.F, or A.B.C.C.C.C.E.F *) + +let intersections d1 d2 = + let rec aux d1 = + if DirPath.is_empty d1 then [d2] else + let rest = aux (snd (chop_dirpath 1 d1)) in + if is_dirpath_prefix_of d1 d2 then drop_dirpath_prefix d1 d2 :: rest + else rest in + aux d1 + +let expand p dir = + let ph = extend_path_with_dirpath p.path_physical dir in + let log = append_dirpath p.path_logical dir in + (ph, log) + +let expand_path dir = + let rec aux = function + | [] -> [] + | p :: l -> + match p.path_type with + | ImplicitPath -> expand p dir :: aux l + | ImplicitRootPath -> + let inters = intersections p.path_logical dir in + List.map (expand p) inters @ aux l + | RootPath -> + if is_dirpath_prefix_of p.path_logical dir then + expand p (drop_dirpath_prefix p.path_logical dir) :: aux l + else aux l in + aux !load_paths diff --git a/library/loadpath.mli b/library/loadpath.mli new file mode 100644 index 00000000..62dc5d59 --- /dev/null +++ b/library/loadpath.mli @@ -0,0 +1,58 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* CUnix.physical_path +(** Get the physical path (filesystem location) of a loadpath. *) + +val logical : t -> DirPath.t +(** Get the logical path (Coq module hierarchy) of a loadpath. *) + +val get_load_paths : unit -> t list +(** Get the current loadpath association. *) + +val get_paths : unit -> CUnix.physical_path list +(** Same as [get_load_paths] but only get the physical part. *) + +val add_load_path : CUnix.physical_path -> path_type -> DirPath.t -> unit +(** [add_load_path phys type log] adds the binding [phys := log] to the current + loadpaths. *) + +val remove_load_path : CUnix.physical_path -> unit +(** Remove the current logical path binding associated to a given physical path, + if any. *) + +val find_load_path : CUnix.physical_path -> t +(** Get the binding associated to a physical path. Raises [Not_found] if there + is none. *) + +val is_in_load_paths : CUnix.physical_path -> bool +(** Whether a physical path is currently bound. *) + +val expand_path : DirPath.t -> (CUnix.physical_path * DirPath.t) list +(** Given a relative logical path, associate the list of absolute physical and + logical paths which are possible expansions of it. *) + +val expand_root_path : DirPath.t -> CUnix.physical_path list +(** As [expand_path] but restricts to root loadpaths. *) diff --git a/library/nameops.ml b/library/nameops.ml index e733d19d..02b085a7 100644 --- a/library/nameops.ml +++ b/library/nameops.ml @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* str "_" @@ -24,20 +24,20 @@ let code_of_0 = Char.code '0' let code_of_9 = Char.code '9' let cut_ident skip_quote s = - let s = string_of_id s in + let s = Id.to_string s in let slen = String.length s in (* [n'] is the position of the first non nullary digit *) let rec numpart n n' = - if n = 0 then + if Int.equal n 0 then (* ident made of _ and digits only [and ' if skip_quote]: don't cut it *) slen else let c = Char.code (String.get s (n-1)) in - if c = code_of_0 && n <> slen then + if Int.equal c code_of_0 && not (Int.equal n slen) then numpart (n-1) n' else if code_of_0 <= c && c <= code_of_9 then numpart (n-1) (n-1) - else if skip_quote & (c = Char.code '\'' || c = Char.code '_') then + else if skip_quote && (Int.equal c (Char.code '\'') || Int.equal c (Char.code '_')) then numpart (n-1) (n-1) else n' @@ -46,9 +46,9 @@ let cut_ident skip_quote s = let repr_ident s = let numstart = cut_ident false s in - let s = string_of_id s in + let s = Id.to_string s in let slen = String.length s in - if numstart = slen then + if Int.equal numstart slen then (s, None) else (String.sub s 0 numstart, @@ -58,24 +58,24 @@ let make_ident sa = function | Some n -> let c = Char.code (String.get sa (String.length sa -1)) in let s = - if c < code_of_0 or c > code_of_9 then sa ^ (string_of_int n) + if c < code_of_0 || c > code_of_9 then sa ^ (string_of_int n) else sa ^ "_" ^ (string_of_int n) in - id_of_string s - | None -> id_of_string (String.copy sa) + Id.of_string s + | None -> Id.of_string (String.copy sa) let root_of_id id = let suffixstart = cut_ident true id in - id_of_string (String.sub (string_of_id id) 0 suffixstart) + Id.of_string (String.sub (Id.to_string id) 0 suffixstart) (* Rem: semantics is a bit different, if an ident starts with toto00 then after successive renamings it comes to toto09, then it goes on with toto10 *) let lift_subscript id = - let id = string_of_id id in + let id = Id.to_string id in let len = String.length id in let rec add carrypos = let c = id.[carrypos] in if is_digit c then - if c = '9' then begin + if Int.equal (Char.code c) (Char.code '9') then begin assert (carrypos>0); add (carrypos-1) end @@ -93,20 +93,20 @@ let lift_subscript id = end; newid end - in id_of_string (add (len-1)) + in Id.of_string (add (len-1)) let has_subscript id = - let id = string_of_id id in + let id = Id.to_string id in is_digit (id.[String.length id - 1]) let forget_subscript id = let numstart = cut_ident false id in let newid = String.make (numstart+1) '0' in - String.blit (string_of_id id) 0 newid 0 numstart; - (id_of_string newid) + String.blit (Id.to_string id) 0 newid 0 numstart; + (Id.of_string newid) -let add_suffix id s = id_of_string (string_of_id id ^ s) -let add_prefix s id = id_of_string (s ^ string_of_id id) +let add_suffix id s = Id.of_string (Id.to_string id ^ s) +let add_prefix s id = Id.of_string (s ^ Id.to_string id) let atompart_of_id id = fst (repr_ident id) @@ -114,7 +114,7 @@ let atompart_of_id id = fst (repr_ident id) let out_name = function | Name id -> id - | Anonymous -> failwith "out_name: expects a defined name" + | Anonymous -> failwith "Nameops.out_name" let name_fold f na a = match na with @@ -136,13 +136,14 @@ let name_fold_map f e = function | Name id -> let (e,id) = f e id in (e,Name id) | Anonymous -> e,Anonymous -let pr_lab l = str (string_of_label l) +let pr_lab l = str (Label.to_string l) -let default_library = Names.initial_dir (* = ["Top"] *) +let default_library = Names.DirPath.initial (* = ["Top"] *) (*s Roots of the space of absolute names *) -let coq_root = id_of_string "Coq" -let default_root_prefix = make_dirpath [] +let coq_string = "Coq" +let coq_root = Id.of_string coq_string +let default_root_prefix = DirPath.empty (* Metavariables *) let pr_meta = Pp.int diff --git a/library/nameops.mli b/library/nameops.mli index 9571d2a3..23432ae2 100644 --- a/library/nameops.mli +++ b/library/nameops.mli @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* Pp.std_ppcmds -val pr_name : name -> Pp.std_ppcmds +val pr_id : Id.t -> Pp.std_ppcmds +val pr_name : Name.t -> Pp.std_ppcmds -val make_ident : string -> int option -> identifier -val repr_ident : identifier -> string * int option +val make_ident : string -> int option -> Id.t +val repr_ident : Id.t -> string * int option -val atompart_of_id : identifier -> string (** remove trailing digits *) -val root_of_id : identifier -> identifier (** remove trailing digits, ' and _ *) +val atompart_of_id : Id.t -> string (** remove trailing digits *) +val root_of_id : Id.t -> Id.t (** remove trailing digits, ' and _ *) -val add_suffix : identifier -> string -> identifier -val add_prefix : string -> identifier -> identifier +val add_suffix : Id.t -> string -> Id.t +val add_prefix : string -> Id.t -> Id.t -val has_subscript : identifier -> bool -val lift_subscript : identifier -> identifier -val forget_subscript : identifier -> identifier +val has_subscript : Id.t -> bool +val lift_subscript : Id.t -> Id.t +val forget_subscript : Id.t -> Id.t -val out_name : name -> identifier +val out_name : Name.t -> Id.t +(** [out_name] associates [id] to [Name id]. Raises [Failure "Nameops.out_name"] + otherwise. *) -val name_fold : (identifier -> 'a -> 'a) -> name -> 'a -> 'a -val name_iter : (identifier -> unit) -> name -> unit -val name_cons : name -> identifier list -> identifier list -val name_app : (identifier -> identifier) -> name -> name -val name_fold_map : ('a -> identifier -> 'a * identifier) -> 'a -> name -> 'a * name +val name_fold : (Id.t -> 'a -> 'a) -> Name.t -> 'a -> 'a +val name_iter : (Id.t -> unit) -> Name.t -> unit +val name_cons : Name.t -> Id.t list -> Id.t list +val name_app : (Id.t -> Id.t) -> Name.t -> Name.t +val name_fold_map : ('a -> Id.t -> 'a * Id.t) -> 'a -> Name.t -> 'a * Name.t -val pr_lab : label -> Pp.std_ppcmds +val pr_lab : Label.t -> Pp.std_ppcmds (** some preset paths *) -val default_library : dir_path +val default_library : DirPath.t (** This is the root of the standard library of Coq *) -val coq_root : module_ident +val coq_root : module_ident (** "Coq" *) +val coq_string : string (** "Coq" *) (** This is the default root prefix for developments which doesn't mention a root *) -val default_root_prefix : dir_path +val default_root_prefix : DirPath.t (** Metavariables *) val pr_meta : Term.metavariable -> Pp.std_ppcmds diff --git a/library/nametab.ml b/library/nametab.ml index 9e8b22b0..6af1e686 100644 --- a/library/nametab.ml +++ b/library/nametab.ml @@ -1,29 +1,24 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* t -> bool val to_string : t -> string - val repr : t -> identifier * module_ident list + val repr : t -> Id.t * module_ident list end +module type EqualityType = +sig + type t + val equal : t -> t -> bool +end (* A ['a t] is a map from [user_name] to ['a], with possible lookup by partially qualified names of type [qualid]. The mapping of @@ -62,68 +63,76 @@ end the same object. *) module type NAMETREE = sig - type 'a t + type elt + type t type user_name - val empty : 'a t - val push : visibility -> user_name -> 'a -> 'a t -> 'a t - val locate : qualid -> 'a t -> 'a - val find : user_name -> 'a t -> 'a - val exists : user_name -> 'a t -> bool - val user_name : qualid -> 'a t -> user_name - val shortest_qualid : Idset.t -> user_name -> 'a t -> qualid - val find_prefixes : qualid -> 'a t -> 'a list + val empty : t + val push : visibility -> user_name -> elt -> t -> t + val locate : qualid -> t -> elt + val find : user_name -> t -> elt + val exists : user_name -> t -> bool + val user_name : qualid -> t -> user_name + val shortest_qualid : Id.Set.t -> user_name -> t -> qualid + val find_prefixes : qualid -> t -> elt list end -module Make(U:UserName) : NAMETREE with type user_name = U.t - = +module Make (U : UserName) (E : EqualityType) : NAMETREE + with type user_name = U.t and type elt = E.t = struct + type elt = E.t type user_name = U.t - type 'a path_status = + type path_status = Nothing - | Relative of user_name * 'a - | Absolute of user_name * 'a + | Relative of user_name * elt + | Absolute of user_name * elt (* Dictionaries of short names *) - type 'a nametree = ('a path_status * 'a nametree ModIdmap.t) + type nametree = + { path : path_status; + map : nametree ModIdmap.t } - type 'a t = 'a nametree Idmap.t + let mktree p m = { path=p; map=m } + let empty_tree = mktree Nothing ModIdmap.empty - let empty = Idmap.empty + type t = nametree Id.Map.t + + let empty = Id.Map.empty (* [push_until] is used to register [Until vis] visibility and [push_exactly] to [Exactly vis] and [push_tree] chooses the right one*) - let rec push_until uname o level (current,dirmap) = function + let rec push_until uname o level tree = function | modid :: path -> - let mc = - try ModIdmap.find modid dirmap - with Not_found -> (Nothing, ModIdmap.empty) - in + let modify _ mc = push_until uname o (level-1) mc path in + let map = + try ModIdmap.modify modid modify tree.map + with Not_found -> + let ptab = modify () empty_tree in + ModIdmap.add modid ptab tree.map + in let this = if level <= 0 then - match current with + match tree.path with | Absolute (n,_) -> (* This is an absolute name, we must keep it otherwise it may become unaccessible forever *) - Flags.if_warn msg_warning (str ("Trying to mask the absolute name \"" ^ U.to_string n ^ "\"!")); - current + tree.path | Nothing | Relative _ -> Relative (uname,o) - else current + else tree.path in - let ptab' = push_until uname o (level-1) mc path in - (this, ModIdmap.add modid ptab' dirmap) + mktree this map | [] -> - match current with + match tree.path with | Absolute (uname',o') -> - if o'=o then begin - assert (uname=uname'); - current, dirmap + if E.equal o' o then begin + assert (U.equal uname uname'); + tree (* we are putting the same thing for the second time :) *) end else @@ -133,56 +142,56 @@ struct error ("Cannot mask the absolute name \"" ^ U.to_string uname' ^ "\"!") | Nothing - | Relative _ -> Absolute (uname,o), dirmap - - -let rec push_exactly uname o level (current,dirmap) = function - | modid :: path -> - let mc = - try ModIdmap.find modid dirmap - with Not_found -> (Nothing, ModIdmap.empty) - in - if level = 0 then - let this = - match current with - | Absolute (n,_) -> - (* This is an absolute name, we must keep it - otherwise it may become unaccessible forever *) - Flags.if_warn - msg_warning (str ("Trying to mask the absolute name \"" - ^ U.to_string n ^ "\"!")); - current - | Nothing - | Relative _ -> Relative (uname,o) - in - (this, dirmap) - else (* not right level *) - let ptab' = push_exactly uname o (level-1) mc path in - (current, ModIdmap.add modid ptab' dirmap) - | [] -> - anomaly "Prefix longer than path! Impossible!" + | Relative _ -> mktree (Absolute (uname,o)) tree.map + + +let rec push_exactly uname o level tree = function +| [] -> + anomaly (Pp.str "Prefix longer than path! Impossible!") +| modid :: path -> + if Int.equal level 0 then + let this = + match tree.path with + | Absolute (n,_) -> + (* This is an absolute name, we must keep it + otherwise it may become unaccessible forever *) + msg_warning (str ("Trying to mask the absolute name \"" + ^ U.to_string n ^ "\"!")); + tree.path + | Nothing + | Relative _ -> Relative (uname,o) + in + mktree this tree.map + else (* not right level *) + let modify _ mc = push_exactly uname o (level-1) mc path in + let map = + try ModIdmap.modify modid modify tree.map + with Not_found -> + let ptab = modify () empty_tree in + ModIdmap.add modid ptab tree.map + in + mktree tree.path map let push visibility uname o tab = let id,dir = U.repr uname in - let ptab = - try Idmap.find id tab - with Not_found -> (Nothing, ModIdmap.empty) - in - let ptab' = match visibility with + let modify _ ptab = match visibility with | Until i -> push_until uname o (i-1) ptab dir | Exactly i -> push_exactly uname o (i-1) ptab dir in - Idmap.add id ptab' tab + try Id.Map.modify id modify tab + with Not_found -> + let ptab = modify () empty_tree in + Id.Map.add id ptab tab -let rec search (current,modidtab) = function - | modid :: path -> search (ModIdmap.find modid modidtab) path - | [] -> current +let rec search tree = function + | modid :: path -> search (ModIdmap.find modid tree.map) path + | [] -> tree.path let find_node qid tab = let (dir,id) = repr_qualid qid in - search (Idmap.find id tab) (repr_dirpath dir) + search (Id.Map.find id tab) (DirPath.repr dir) let locate qid tab = let o = match find_node qid tab with @@ -200,7 +209,7 @@ let user_name qid tab = let find uname tab = let id,l = U.repr uname in - match search (Idmap.find id tab) l with + match search (Id.Map.find id tab) l with Absolute (_,o) -> o | _ -> raise Not_found @@ -213,36 +222,38 @@ let exists uname tab = let shortest_qualid ctx uname tab = let id,dir = U.repr uname in - let hidden = Idset.mem id ctx in - let rec find_uname pos dir (path,tab) = match path with + let hidden = Id.Set.mem id ctx in + let rec find_uname pos dir tree = + let is_empty = match pos with [] -> true | _ -> false in + match tree.path with | Absolute (u,_) | Relative (u,_) - when u=uname && not(pos=[] && hidden) -> List.rev pos + when U.equal u uname && not (is_empty && hidden) -> List.rev pos | _ -> match dir with [] -> raise Not_found - | id::dir -> find_uname (id::pos) dir (ModIdmap.find id tab) + | id::dir -> find_uname (id::pos) dir (ModIdmap.find id tree.map) in - let ptab = Idmap.find id tab in + let ptab = Id.Map.find id tab in let found_dir = find_uname [] dir ptab in - make_qualid (make_dirpath found_dir) id + make_qualid (DirPath.make found_dir) id let push_node node l = match node with - | Absolute (_,o) | Relative (_,o) when not (List.mem o l) -> o::l + | Absolute (_,o) | Relative (_,o) when not (List.mem_f E.equal o l) -> o::l | _ -> l let rec flatten_idmap tab l = - ModIdmap.fold (fun _ (current,idtab) l -> - flatten_idmap idtab (push_node current l)) tab l + let f _ tree l = flatten_idmap tree.map (push_node tree.path l) in + ModIdmap.fold f tab l -let rec search_prefixes (current,modidtab) = function - | modid :: path -> search_prefixes (ModIdmap.find modid modidtab) path - | [] -> List.rev (flatten_idmap modidtab (push_node current [])) +let rec search_prefixes tree = function + | modid :: path -> search_prefixes (ModIdmap.find modid tree.map) path + | [] -> List.rev (flatten_idmap tree.map (push_node tree.path [])) let find_prefixes qid tab = try let (dir,id) = repr_qualid qid in - search_prefixes (Idmap.find id tab) (repr_dirpath dir) + search_prefixes (Id.Map.find id tab) (DirPath.repr dir) with Not_found -> [] end @@ -251,49 +262,65 @@ end (* Global name tables *************************************************) -module SpTab = Make (struct - type t = full_path - let to_string = string_of_path - let repr sp = - let dir,id = repr_path sp in - id, (repr_dirpath dir) - end) +module FullPath = +struct + type t = full_path + let equal = eq_full_path + let to_string = string_of_path + let repr sp = + let dir,id = repr_path sp in + id, (DirPath.repr dir) +end + +module ExtRefEqual = ExtRefOrdered +module KnEqual = Names.KerName +module MPEqual = Names.ModPath +module ExtRefTab = Make(FullPath)(ExtRefEqual) +module KnTab = Make(FullPath)(KnEqual) +module MPTab = Make(FullPath)(MPEqual) -type ccitab = extended_global_reference SpTab.t -let the_ccitab = ref (SpTab.empty : ccitab) +type ccitab = ExtRefTab.t +let the_ccitab = ref (ExtRefTab.empty : ccitab) -type kntab = kernel_name SpTab.t -let the_tactictab = ref (SpTab.empty : kntab) +type kntab = KnTab.t +let the_tactictab = ref (KnTab.empty : kntab) -type mptab = module_path SpTab.t -let the_modtypetab = ref (SpTab.empty : mptab) +type mptab = MPTab.t +let the_modtypetab = ref (MPTab.empty : mptab) + +module DirPath' = +struct + include DirPath + let repr dir = match DirPath.repr dir with + | [] -> anomaly (Pp.str "Empty dirpath") + | id :: l -> (id, l) +end +module GlobDir = +struct + type t = global_dir_reference + let equal = eq_global_dir_reference +end -module DirTab = Make(struct - type t = dir_path - let to_string = string_of_dirpath - let repr dir = match repr_dirpath dir with - | [] -> anomaly "Empty dirpath" - | id::l -> (id,l) - end) +module DirTab = Make(DirPath')(GlobDir) (* If we have a (closed) module M having a submodule N, than N does not have the entry in [the_dirtab]. *) -type dirtab = global_dir_reference DirTab.t +type dirtab = DirTab.t let the_dirtab = ref (DirTab.empty : dirtab) (* Reversed name tables ***************************************************) (* This table translates extended_global_references back to section paths *) -module Globrevtab = Map.Make(ExtRefOrdered) +module Globrevtab = HMap.Make(ExtRefOrdered) type globrevtab = full_path Globrevtab.t let the_globrevtab = ref (Globrevtab.empty : globrevtab) -type mprevtab = dir_path MPmap.t +type mprevtab = DirPath.t MPmap.t let the_modrevtab = ref (MPmap.empty : mprevtab) type mptrevtab = full_path MPmap.t @@ -312,19 +339,19 @@ let the_tacticrevtab = ref (KNmap.empty : knrevtab) let push_xref visibility sp xref = match visibility with | Until _ -> - the_ccitab := SpTab.push visibility sp xref !the_ccitab; + the_ccitab := ExtRefTab.push visibility sp xref !the_ccitab; the_globrevtab := Globrevtab.add xref sp !the_globrevtab | _ -> begin - if SpTab.exists sp !the_ccitab then - match SpTab.find sp !the_ccitab with + if ExtRefTab.exists sp !the_ccitab then + match ExtRefTab.find sp !the_ccitab with | TrueGlobal( ConstRef _) | TrueGlobal( IndRef _) | TrueGlobal( ConstructRef _) as xref -> - the_ccitab := SpTab.push visibility sp xref !the_ccitab; + the_ccitab := ExtRefTab.push visibility sp xref !the_ccitab; | _ -> - the_ccitab := SpTab.push visibility sp xref !the_ccitab; + the_ccitab := ExtRefTab.push visibility sp xref !the_ccitab; else - the_ccitab := SpTab.push visibility sp xref !the_ccitab; + the_ccitab := ExtRefTab.push visibility sp xref !the_ccitab; end let push_cci visibility sp ref = @@ -337,13 +364,13 @@ let push_syndef visibility sp kn = let push = push_cci let push_modtype vis sp kn = - the_modtypetab := SpTab.push vis sp kn !the_modtypetab; + the_modtypetab := MPTab.push vis sp kn !the_modtypetab; the_modtyperevtab := MPmap.add kn sp !the_modtyperevtab (* This is for tactic definition names *) let push_tactic vis sp kn = - the_tactictab := SpTab.push vis sp kn !the_tactictab; + the_tactictab := KnTab.push vis sp kn !the_tactictab; the_tacticrevtab := KNmap.add kn sp !the_tacticrevtab @@ -359,22 +386,22 @@ let push_dir vis dir dir_ref = (* This should be used when syntactic definitions are allowed *) -let locate_extended qid = SpTab.locate qid !the_ccitab +let locate_extended qid = ExtRefTab.locate qid !the_ccitab (* This should be used when no syntactic definitions is expected *) let locate qid = match locate_extended qid with | TrueGlobal ref -> ref | SynDef _ -> raise Not_found -let full_name_cci qid = SpTab.user_name qid !the_ccitab +let full_name_cci qid = ExtRefTab.user_name qid !the_ccitab let locate_syndef qid = match locate_extended qid with | TrueGlobal _ -> raise Not_found | SynDef kn -> kn -let locate_modtype qid = SpTab.locate qid !the_modtypetab -let full_name_modtype qid = SpTab.user_name qid !the_modtypetab +let locate_modtype qid = MPTab.locate qid !the_modtypetab +let full_name_modtype qid = MPTab.user_name qid !the_modtypetab -let locate_tactic qid = SpTab.locate qid !the_tactictab +let locate_tactic qid = KnTab.locate qid !the_tactictab let locate_dir qid = DirTab.locate qid !the_dirtab @@ -396,9 +423,15 @@ let locate_section qid = let locate_all qid = List.fold_right (fun a l -> match a with TrueGlobal a -> a::l | _ -> l) - (SpTab.find_prefixes qid !the_ccitab) [] + (ExtRefTab.find_prefixes qid !the_ccitab) [] + +let locate_extended_all qid = ExtRefTab.find_prefixes qid !the_ccitab -let locate_extended_all qid = SpTab.find_prefixes qid !the_ccitab +let locate_extended_all_tactic qid = KnTab.find_prefixes qid !the_tactictab + +let locate_extended_all_dir qid = DirTab.find_prefixes qid !the_dirtab + +let locate_extended_all_modtype qid = MPTab.find_prefixes qid !the_modtypetab (* Derived functions *) @@ -408,11 +441,11 @@ let locate_constant qid = | _ -> raise Not_found let global_of_path sp = - match SpTab.find sp !the_ccitab with + match ExtRefTab.find sp !the_ccitab with | TrueGlobal ref -> ref | _ -> raise Not_found -let extended_global_of_path sp = SpTab.find sp !the_ccitab +let extended_global_of_path sp = ExtRefTab.find sp !the_ccitab let global r = let (loc,qid) = qualid_of_reference r in @@ -427,7 +460,7 @@ let global r = (* Exists functions ********************************************************) -let exists_cci sp = SpTab.exists sp !the_ccitab +let exists_cci sp = ExtRefTab.exists sp !the_ccitab let exists_dir dir = DirTab.exists dir !the_dirtab @@ -435,13 +468,15 @@ let exists_section = exists_dir let exists_module = exists_dir -let exists_modtype sp = SpTab.exists sp !the_modtypetab +let exists_modtype sp = MPTab.exists sp !the_modtypetab + +let exists_tactic kn = KnTab.exists kn !the_tactictab (* Reverse locate functions ***********************************************) let path_of_global ref = match ref with - | VarRef id -> make_path empty_dirpath id + | VarRef id -> make_path DirPath.empty id | _ -> Globrevtab.find (TrueGlobal ref) !the_globrevtab let dirpath_of_global ref = @@ -459,37 +494,37 @@ let dirpath_of_module mp = let path_of_tactic kn = KNmap.find kn !the_tacticrevtab +let path_of_modtype mp = + MPmap.find mp !the_modtyperevtab + (* Shortest qualid functions **********************************************) let shortest_qualid_of_global ctx ref = match ref with - | VarRef id -> make_qualid empty_dirpath id + | VarRef id -> make_qualid DirPath.empty id | _ -> let sp = Globrevtab.find (TrueGlobal ref) !the_globrevtab in - SpTab.shortest_qualid ctx sp !the_ccitab + ExtRefTab.shortest_qualid ctx sp !the_ccitab let shortest_qualid_of_syndef ctx kn = let sp = path_of_syndef kn in - SpTab.shortest_qualid ctx sp !the_ccitab + ExtRefTab.shortest_qualid ctx sp !the_ccitab let shortest_qualid_of_module mp = let dir = MPmap.find mp !the_modrevtab in - DirTab.shortest_qualid Idset.empty dir !the_dirtab + DirTab.shortest_qualid Id.Set.empty dir !the_dirtab let shortest_qualid_of_modtype kn = let sp = MPmap.find kn !the_modtyperevtab in - SpTab.shortest_qualid Idset.empty sp !the_modtypetab + MPTab.shortest_qualid Id.Set.empty sp !the_modtypetab let shortest_qualid_of_tactic kn = let sp = KNmap.find kn !the_tacticrevtab in - SpTab.shortest_qualid Idset.empty sp !the_tactictab + KnTab.shortest_qualid Id.Set.empty sp !the_tactictab let pr_global_env env ref = - (* Il est important de laisser le let-in, car les streams s'évaluent - paresseusement : il faut forcer l'évaluation pour capturer - l'éventuelle levée d'une exception (le cas échoit dans le debugger) *) - let s = string_of_qualid (shortest_qualid_of_global env ref) in - (str s) + try str (string_of_qualid (shortest_qualid_of_global env ref)) + with Not_found as e -> prerr_endline "pr_global_env not found"; raise e let global_inductive r = match global r with @@ -504,22 +539,10 @@ let global_inductive r = (********************************************************************) (* Registration of tables as a global table and rollback *) -type frozen = ccitab * dirtab * kntab * kntab - * globrevtab * mprevtab * knrevtab * knrevtab - -let init () = - the_ccitab := SpTab.empty; - the_dirtab := DirTab.empty; - the_modtypetab := SpTab.empty; - the_tactictab := SpTab.empty; - the_globrevtab := Globrevtab.empty; - the_modrevtab := MPmap.empty; - the_modtyperevtab := MPmap.empty; - the_tacticrevtab := KNmap.empty - - +type frozen = ccitab * dirtab * mptab * kntab + * globrevtab * mprevtab * mptrevtab * knrevtab -let freeze () = +let freeze _ : frozen = !the_ccitab, !the_dirtab, !the_modtypetab, @@ -543,7 +566,7 @@ let _ = Summary.declare_summary "names" { Summary.freeze_function = freeze; Summary.unfreeze_function = unfreeze; - Summary.init_function = init } + Summary.init_function = Summary.nop } (* Deprecated synonyms *) diff --git a/library/nametab.mli b/library/nametab.mli index 79ea119b..e3aeb675 100644 --- a/library/nametab.mli +++ b/library/nametab.mli @@ -1,15 +1,15 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* bool] @@ -51,19 +51,17 @@ open Libnames {- [shortest_qualid_of : object_reference -> user_name] The [user_name] can be for example the shortest non ambiguous [qualid] or - the [full_user_name] or [identifier]. Such a function can also have a + the [full_user_name] or [Id.t]. Such a function can also have a local context argument.}} *) exception GlobalizationError of qualid -exception GlobalizationConstantError of qualid (** Raises a globalization error *) -val error_global_not_found_loc : loc -> qualid -> 'a +val error_global_not_found_loc : Loc.t -> qualid -> 'a val error_global_not_found : qualid -> 'a -val error_global_constant_not_found_loc : loc -> qualid -> 'a (** {6 Register visibility of things } *) @@ -79,7 +77,7 @@ type visibility = Until of int | Exactly of int val push : visibility -> full_path -> global_reference -> unit val push_modtype : visibility -> full_path -> module_path -> unit -val push_dir : visibility -> dir_path -> global_dir_reference -> unit +val push_dir : visibility -> DirPath.t -> global_dir_reference -> unit val push_syndef : visibility -> full_path -> syndef_name -> unit type ltac_constant = kernel_name @@ -98,7 +96,7 @@ val locate_syndef : qualid -> syndef_name val locate_modtype : qualid -> module_path val locate_dir : qualid -> global_dir_reference val locate_module : qualid -> module_path -val locate_section : qualid -> dir_path +val locate_section : qualid -> DirPath.t val locate_tactic : qualid -> ltac_constant (** These functions globalize user-level references into global @@ -113,6 +111,9 @@ val global_inductive : reference -> inductive val locate_all : qualid -> global_reference list val locate_extended_all : qualid -> extended_global_reference list +val locate_extended_all_tactic : qualid -> ltac_constant list +val locate_extended_all_dir : qualid -> global_dir_reference list +val locate_extended_all_modtype : qualid -> module_path list (** Mapping a full path to a global reference *) @@ -123,15 +124,16 @@ val extended_global_of_path : full_path -> extended_global_reference val exists_cci : full_path -> bool val exists_modtype : full_path -> bool -val exists_dir : dir_path -> bool -val exists_section : dir_path -> bool (** deprecated synonym of [exists_dir] *) -val exists_module : dir_path -> bool (** deprecated synonym of [exists_dir] *) +val exists_dir : DirPath.t -> bool +val exists_section : DirPath.t -> bool (** deprecated synonym of [exists_dir] *) +val exists_module : DirPath.t -> bool (** deprecated synonym of [exists_dir] *) +val exists_tactic : full_path -> bool (** deprecated synonym of [exists_dir] *) (** {6 These functions locate qualids into full user names } *) val full_name_cci : qualid -> full_path val full_name_modtype : qualid -> full_path -val full_name_module : qualid -> dir_path +val full_name_module : qualid -> DirPath.t (** {6 Reverse lookup } Finding user names corresponding to the given @@ -142,25 +144,28 @@ val full_name_module : qualid -> dir_path val path_of_syndef : syndef_name -> full_path val path_of_global : global_reference -> full_path -val dirpath_of_module : module_path -> dir_path +val dirpath_of_module : module_path -> DirPath.t +val path_of_modtype : module_path -> full_path val path_of_tactic : ltac_constant -> full_path (** Returns in particular the dirpath or the basename of the full path associated to global reference *) -val dirpath_of_global : global_reference -> dir_path -val basename_of_global : global_reference -> identifier +val dirpath_of_global : global_reference -> DirPath.t +val basename_of_global : global_reference -> Id.t -(** Printing of global references using names as short as possible *) -val pr_global_env : Idset.t -> global_reference -> std_ppcmds +(** Printing of global references using names as short as possible. + @raise Not_found when the reference is not in the global tables. *) +val pr_global_env : Id.Set.t -> global_reference -> std_ppcmds (** The [shortest_qualid] functions given an object with [user_name] Coq.A.B.x, try to find the shortest among x, B.x, A.B.x and - Coq.A.B.x that denotes the same object. *) + Coq.A.B.x that denotes the same object. + @raise Not_found for unknown objects. *) -val shortest_qualid_of_global : Idset.t -> global_reference -> qualid -val shortest_qualid_of_syndef : Idset.t -> syndef_name -> qualid +val shortest_qualid_of_global : Id.Set.t -> global_reference -> qualid +val shortest_qualid_of_syndef : Id.Set.t -> syndef_name -> qualid val shortest_qualid_of_modtype : module_path -> qualid val shortest_qualid_of_module : module_path -> qualid val shortest_qualid_of_tactic : ltac_constant -> qualid diff --git a/library/states.ml b/library/states.ml index 768fbb23..a1c2a095 100644 --- a/library/states.ml +++ b/library/states.ml @@ -1,46 +1,46 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* - if !Flags.load_proofs <> Flags.Force then - Util.error "Write State only works with option -force-load-proofs"; - raw_extern s (freeze())), + let s = ensure_suffix s in + raw_extern s (freeze ~marshallable:`Yes)), (fun s -> - unfreeze - (with_magic_number_check (raw_intern (Library.get_load_paths ())) s); + let s = ensure_suffix s in + let paths = Loadpath.get_paths () in + unfreeze (with_magic_number_check (raw_intern paths) s); Library.overwrite_library_filenames s) (* Rollback. *) -let with_heavy_rollback f h x = - let st = freeze () in - try - f x - with reraise -> - let e = h reraise in (unfreeze st; raise e) - let with_state_protection f x = - let st = freeze () in + let st = freeze ~marshallable:`No in try let a = f x in unfreeze st; a with reraise -> - (unfreeze st; raise reraise) + let reraise = Errors.push reraise in + (unfreeze st; iraise reraise) + +let with_state_protection_on_exception = Future.transactify diff --git a/library/states.mli b/library/states.mli index 9474d831..66de1490 100644 --- a/library/states.mli +++ b/library/states.mli @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* unit val extern_state : string -> unit type state -val freeze : unit -> state +val freeze : marshallable:Summary.marshallable -> state val unfreeze : state -> unit -(** {6 Rollback } *) +val summary_of_state : state -> Summary.frozen -(** [with_heavy_rollback f x] applies [f] to [x] and restores the - state of the whole system as it was before the evaluation if an exception - is raised. *) -val with_heavy_rollback : ('a -> 'b) -> (exn -> exn) -> 'a -> 'b +(** {6 Rollback } *) (** [with_state_protection f x] applies [f] to [x] and restores the - state of the whole system as it was before the evaluation of f *) + state of the whole system as it was before applying [f] *) val with_state_protection : ('a -> 'b) -> 'a -> 'b +(** [with_state_protection_on_exception f x] applies [f] to [x] and restores the + state of the whole system as it was before applying [f] only if an + exception is raised. Unlike [with_state_protection] it also takes into + account the proof state *) + +val with_state_protection_on_exception : ('a -> 'b) -> 'a -> 'b diff --git a/library/summary.ml b/library/summary.ml index 4e7a1f8e..7e7628a1 100644 --- a/library/summary.ml +++ b/library/summary.ml @@ -1,25 +1,28 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* 'a; + freeze_function : marshallable -> 'a; unfreeze_function : 'a -> unit; init_function : unit -> unit } -let summaries = - (Hashtbl.create 17 : (string, Dyn.t summary_declaration) Hashtbl.t) +let summaries = ref Int.Map.empty -let internal_declare_summary sumname sdecl = - let (infun,outfun) = Dyn.create sumname in - let dyn_freeze () = infun (sdecl.freeze_function()) +let mangle id = id ^ "-SUMMARY" + +let internal_declare_summary hash sumname sdecl = + let (infun, outfun) = Dyn.create (mangle sumname) in + let dyn_freeze b = infun (sdecl.freeze_function b) and dyn_unfreeze sum = sdecl.unfreeze_function (outfun sum) and dyn_init = sdecl.init_function in let ddecl = { @@ -27,30 +30,148 @@ let internal_declare_summary sumname sdecl = unfreeze_function = dyn_unfreeze; init_function = dyn_init } in - if Hashtbl.mem summaries sumname then - anomalylabstrm "Summary.declare_summary" - (str "Cannot declare a summary twice: " ++ str sumname); - Hashtbl.add summaries sumname ddecl + summaries := Int.Map.add hash (sumname, ddecl) !summaries + +let all_declared_summaries = ref Int.Set.empty + +let summary_names = ref [] +let name_of_summary name = + try List.assoc name !summary_names + with Not_found -> "summary name not found" let declare_summary sumname decl = - internal_declare_summary (sumname^"-SUMMARY") decl + let hash = String.hash sumname in + let () = if Int.Map.mem hash !summaries then + let (name, _) = Int.Map.find hash !summaries in + anomaly ~label:"Summary.declare_summary" + (str "Colliding summary names: " ++ str sumname ++ str " vs. " ++ str name) + in + all_declared_summaries := Int.Set.add hash !all_declared_summaries; + summary_names := (hash, sumname) :: !summary_names; + internal_declare_summary hash sumname decl + +type frozen = { + summaries : (int * Dyn.t) list; + (** Ordered list w.r.t. the first component. *) + ml_module : Dyn.t option; + (** Special handling of the ml_module summary. *) +} -type frozen = Dyn.t Stringmap.t +let empty_frozen = { summaries = []; ml_module = None; } -let freeze_summaries () = - let m = ref Stringmap.empty in - Hashtbl.iter - (fun id decl -> m := Stringmap.add id (decl.freeze_function()) !m) - summaries; - !m +let ml_modules = "ML-MODULES" +let ml_modules_summary = String.hash ml_modules +let freeze_summaries ~marshallable : frozen = + let fold id (_, decl) accu = + (* to debug missing Lazy.force + if marshallable <> `No then begin + prerr_endline ("begin marshalling " ^ id); + ignore(Marshal.to_string (decl.freeze_function marshallable) []); + prerr_endline ("end marshalling " ^ id); + end; + /debug *) + let state = decl.freeze_function marshallable in + if Int.equal id ml_modules_summary then { accu with ml_module = Some state } + else { accu with summaries = (id, state) :: accu.summaries } + in + Int.Map.fold_right fold !summaries empty_frozen let unfreeze_summaries fs = - Hashtbl.iter - (fun id decl -> - try decl.unfreeze_function (Stringmap.find id fs) - with Not_found -> decl.init_function()) - summaries + (* The unfreezing of [ml_modules_summary] has to be anticipated since it + * may modify the content of [summaries] ny loading new ML modules *) + let (_, decl) = + try Int.Map.find ml_modules_summary !summaries + with Not_found -> anomaly (str "Undeclared summary " ++ str ml_modules) + in + let () = match fs.ml_module with + | None -> anomaly (str "Undeclared summary " ++ str ml_modules) + | Some state -> decl.unfreeze_function state + in + let fold id (_, decl) states = + if Int.equal id ml_modules_summary then states + else match states with + | [] -> + let () = decl.init_function () in + [] + | (nid, state) :: rstates -> + if Int.equal id nid then + let () = decl.unfreeze_function state in rstates + else + let () = decl.init_function () in states + in + let fold id decl state = + try fold id decl state + with e when Errors.noncritical e -> + let e = Errors.push e in + Printf.eprintf "Error unfrezing summay %s\n%s\n%!" + (name_of_summary id) (Pp.string_of_ppcmds (Errors.iprint e)); + iraise e + in + (** We rely on the order of the frozen list, and the order of folding *) + ignore (Int.Map.fold_left fold !summaries fs.summaries) let init_summaries () = - Hashtbl.iter (fun _ decl -> decl.init_function()) summaries + Int.Map.iter (fun _ (_, decl) -> decl.init_function ()) !summaries + +(** For global tables registered statically before the end of coqtop + launch, the following empty [init_function] could be used. *) + +let nop () = () + +(** Selective freeze *) + +type frozen_bits = (int * Dyn.t) list + +let ids_of_string_list complement ids = + if not complement then List.map String.hash ids + else + let fold accu id = + let id = String.hash id in + Int.Set.remove id accu + in + let ids = List.fold_left fold !all_declared_summaries ids in + Int.Set.elements ids + +let freeze_summary ~marshallable ?(complement=false) ids = + let ids = ids_of_string_list complement ids in + List.map (fun id -> + let (_, summary) = Int.Map.find id !summaries in + id, summary.freeze_function marshallable) + ids + +let unfreeze_summary datas = + List.iter + (fun (id, data) -> + let (name, summary) = Int.Map.find id !summaries in + try summary.unfreeze_function data + with e -> + let e = Errors.push e in + prerr_endline ("Exception unfreezing " ^ name); + iraise e) + datas + +let surgery_summary { summaries; ml_module } bits = + let summaries = List.map (fun (id, _ as orig) -> + try id, List.assoc id bits + with Not_found -> orig) + summaries in + { summaries; ml_module } + +let project_summary { summaries; ml_module } ?(complement=false) ids = + let ids = ids_of_string_list complement ids in + List.filter (fun (id, _) -> List.mem id ids) summaries + +let pointer_equal l1 l2 = + CList.for_all2eq + (fun (id1,v1) (id2,v2) -> id1 = id2 && Dyn.pointer_equal v1 v2) l1 l2 + +(** All-in-one reference declaration + registration *) + +let ref ?(freeze=fun _ r -> r) ~name x = + let r = ref x in + declare_summary name + { freeze_function = (fun b -> freeze b !r); + unfreeze_function = ((:=) r); + init_function = (fun () -> r := x) }; + r diff --git a/library/summary.mli b/library/summary.mli index 9705af65..48c9390d 100644 --- a/library/summary.mli +++ b/library/summary.mli @@ -1,6 +1,6 @@ (************************************************************************) (* v * The Coq Proof Assistant / The Coq Development Team *) -(* 'a; + (** freeze_function [true] is for marshalling to disk. + * e.g. lazy must be forced *) + freeze_function : marshallable -> 'a; unfreeze_function : 'a -> unit; init_function : unit -> unit } +(** For tables registered during the launch of coqtop, the [init_function] + will be run only once, during an [init_summaries] done at the end of + coqtop initialization. For tables registered later (for instance + during a plugin dynlink), [init_function] is used when unfreezing + an earlier frozen state that doesn't contain any value for this table. + + Beware: for tables registered dynamically after the initialization + of Coq, their init functions may not be run immediately. It is hence + the responsability of plugins to initialize themselves properly. +*) + val declare_summary : string -> 'a summary_declaration -> unit +(** All-in-one reference declaration + summary registration. + It behaves just as OCaml's standard [ref] function, except + that a [declare_summary] is done, with [name] as string. + The [init_function] restores the reference to its initial value. + The [freeze_function] can be overridden *) + +val ref : ?freeze:(marshallable -> 'a -> 'a) -> name:string -> 'a -> 'a ref + +(** Special name for the summary of ML modules. This summary entry is + special because its unfreeze may load ML code and hence add summary + entries. Thus is has to be recognizable, and handled appropriately *) +val ml_modules : string + +(** For global tables registered statically before the end of coqtop + launch, the following empty [init_function] could be used. *) + +val nop : unit -> unit + +(** The type [frozen] is a snapshot of the states of all the registered + tables of the system. *) + type frozen -val freeze_summaries : unit -> frozen +val empty_frozen : frozen +val freeze_summaries : marshallable:marshallable -> frozen val unfreeze_summaries : frozen -> unit val init_summaries : unit -> unit -(** Beware: if some code is dynamically loaded via dynlink after the - initialization of Coq, the init functions of any summary declared - by this code may not be run. It is hence the responsability of - plugins to initialize themselves properly. -*) +(** The type [frozen_bits] is a snapshot of some of the registered tables *) +type frozen_bits +val freeze_summary : + marshallable:marshallable -> ?complement:bool -> string list -> frozen_bits +val unfreeze_summary : frozen_bits -> unit +val surgery_summary : frozen -> frozen_bits -> frozen +val project_summary : frozen -> ?complement:bool -> string list -> frozen_bits +val pointer_equal : frozen_bits -> frozen_bits -> bool diff --git a/library/universes.ml b/library/universes.ml new file mode 100644 index 00000000..79070763 --- /dev/null +++ b/library/universes.ml @@ -0,0 +1,1006 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* Level.pr l + +type universe_constraint_type = ULe | UEq | ULub + +type universe_constraint = universe * universe_constraint_type * universe + +module Constraints = struct + module S = Set.Make( + struct + type t = universe_constraint + + let compare_type c c' = + match c, c' with + | ULe, ULe -> 0 + | ULe, _ -> -1 + | _, ULe -> 1 + | UEq, UEq -> 0 + | UEq, _ -> -1 + | ULub, ULub -> 0 + | ULub, _ -> 1 + + let compare (u,c,v) (u',c',v') = + let i = compare_type c c' in + if Int.equal i 0 then + let i' = Universe.compare u u' in + if Int.equal i' 0 then Universe.compare v v' + else + if c != ULe && Universe.compare u v' = 0 && Universe.compare v u' = 0 then 0 + else i' + else i + end) + + include S + + let add (l,d,r as cst) s = + if Universe.equal l r then s + else add cst s + + let tr_dir = function + | ULe -> Le + | UEq -> Eq + | ULub -> Eq + + let op_str = function ULe -> " <= " | UEq -> " = " | ULub -> " /\\ " + + let pr c = + fold (fun (u1,op,u2) pp_std -> + pp_std ++ Universe.pr u1 ++ str (op_str op) ++ + Universe.pr u2 ++ fnl ()) c (str "") + + let equal x y = + x == y || equal x y + +end + +type universe_constraints = Constraints.t +type 'a universe_constrained = 'a * universe_constraints + +type 'a universe_constraint_function = 'a -> 'a -> universe_constraints -> universe_constraints + +let enforce_eq_instances_univs strict x y c = + let d = if strict then ULub else UEq in + let ax = Instance.to_array x and ay = Instance.to_array y in + if Array.length ax != Array.length ay then + Errors.anomaly (Pp.str "Invalid argument: enforce_eq_instances_univs called with" ++ + Pp.str " instances of different lengths"); + CArray.fold_right2 + (fun x y -> Constraints.add (Universe.make x, d, Universe.make y)) + ax ay c + +let subst_univs_universe_constraint fn (u,d,v) = + let u' = subst_univs_universe fn u and v' = subst_univs_universe fn v in + if Universe.equal u' v' then None + else Some (u',d,v') + +let subst_univs_universe_constraints subst csts = + Constraints.fold + (fun c -> Option.fold_right Constraints.add (subst_univs_universe_constraint subst c)) + csts Constraints.empty + + +let to_constraints g s = + let tr (x,d,y) acc = + let add l d l' acc = Constraint.add (l,Constraints.tr_dir d,l') acc in + match Universe.level x, d, Universe.level y with + | Some l, (ULe | UEq | ULub), Some l' -> add l d l' acc + | _, ULe, Some l' -> enforce_leq x y acc + | _, ULub, _ -> acc + | _, d, _ -> + let f = if d == ULe then check_leq else check_eq in + if f g x y then acc else + raise (Invalid_argument + "to_constraints: non-trivial algebraic constraint between universes") + in Constraints.fold tr s Constraint.empty + +let eq_constr_univs_infer univs m n = + if m == n then true, Constraints.empty + else + let cstrs = ref Constraints.empty in + let eq_universes strict = Univ.Instance.check_eq univs in + let eq_sorts s1 s2 = + if Sorts.equal s1 s2 then true + else + let u1 = Sorts.univ_of_sort s1 and u2 = Sorts.univ_of_sort s2 in + if Univ.check_eq univs u1 u2 then true + else + (cstrs := Constraints.add (u1, UEq, u2) !cstrs; + true) + in + let rec eq_constr' m n = + m == n || Constr.compare_head_gen eq_universes eq_sorts eq_constr' m n + in + let res = Constr.compare_head_gen eq_universes eq_sorts eq_constr' m n in + res, !cstrs + +let leq_constr_univs_infer univs m n = + if m == n then true, Constraints.empty + else + let cstrs = ref Constraints.empty in + let eq_universes strict l l' = Univ.Instance.check_eq univs l l' in + let eq_sorts s1 s2 = + if Sorts.equal s1 s2 then true + else + let u1 = Sorts.univ_of_sort s1 and u2 = Sorts.univ_of_sort s2 in + if Univ.check_eq univs u1 u2 then true + else (cstrs := Constraints.add (u1, UEq, u2) !cstrs; + true) + in + let leq_sorts s1 s2 = + if Sorts.equal s1 s2 then true + else + let u1 = Sorts.univ_of_sort s1 and u2 = Sorts.univ_of_sort s2 in + if Univ.check_leq univs u1 u2 then true + else + (cstrs := Constraints.add (u1, ULe, u2) !cstrs; + true) + in + let rec eq_constr' m n = + m == n || Constr.compare_head_gen eq_universes eq_sorts eq_constr' m n + in + let rec compare_leq m n = + Constr.compare_head_gen_leq eq_universes eq_sorts leq_sorts + eq_constr' leq_constr' m n + and leq_constr' m n = m == n || compare_leq m n in + let res = compare_leq m n in + res, !cstrs + +let eq_constr_universes m n = + if m == n then true, Constraints.empty + else + let cstrs = ref Constraints.empty in + let eq_universes strict l l' = + cstrs := enforce_eq_instances_univs strict l l' !cstrs; true in + let eq_sorts s1 s2 = + if Sorts.equal s1 s2 then true + else + (cstrs := Constraints.add + (Sorts.univ_of_sort s1, UEq, Sorts.univ_of_sort s2) !cstrs; + true) + in + let rec eq_constr' m n = + m == n || Constr.compare_head_gen eq_universes eq_sorts eq_constr' m n + in + let res = Constr.compare_head_gen eq_universes eq_sorts eq_constr' m n in + res, !cstrs + +let leq_constr_universes m n = + if m == n then true, Constraints.empty + else + let cstrs = ref Constraints.empty in + let eq_universes strict l l' = + cstrs := enforce_eq_instances_univs strict l l' !cstrs; true in + let eq_sorts s1 s2 = + if Sorts.equal s1 s2 then true + else (cstrs := Constraints.add + (Sorts.univ_of_sort s1,UEq,Sorts.univ_of_sort s2) !cstrs; + true) + in + let leq_sorts s1 s2 = + if Sorts.equal s1 s2 then true + else + (cstrs := Constraints.add + (Sorts.univ_of_sort s1,ULe,Sorts.univ_of_sort s2) !cstrs; + true) + in + let rec eq_constr' m n = + m == n || Constr.compare_head_gen eq_universes eq_sorts eq_constr' m n + in + let rec compare_leq m n = + Constr.compare_head_gen_leq eq_universes eq_sorts leq_sorts eq_constr' leq_constr' m n + and leq_constr' m n = m == n || compare_leq m n in + let res = compare_leq m n in + res, !cstrs + +let compare_head_gen_proj env equ eqs eqc' m n = + match kind_of_term m, kind_of_term n with + | Proj (p, c), App (f, args) + | App (f, args), Proj (p, c) -> + (match kind_of_term f with + | Const (p', u) when eq_constant (Projection.constant p) p' -> + let pb = Environ.lookup_projection p env in + let npars = pb.Declarations.proj_npars in + if Array.length args == npars + 1 then + eqc' c args.(npars) + else false + | _ -> false) + | _ -> Constr.compare_head_gen equ eqs eqc' m n + +let eq_constr_universes_proj env m n = + if m == n then true, Constraints.empty + else + let cstrs = ref Constraints.empty in + let eq_universes strict l l' = + cstrs := enforce_eq_instances_univs strict l l' !cstrs; true in + let eq_sorts s1 s2 = + if Sorts.equal s1 s2 then true + else + (cstrs := Constraints.add + (Sorts.univ_of_sort s1, UEq, Sorts.univ_of_sort s2) !cstrs; + true) + in + let rec eq_constr' m n = + m == n || compare_head_gen_proj env eq_universes eq_sorts eq_constr' m n + in + let res = eq_constr' m n in + res, !cstrs + +(* Generator of levels *) +let new_univ_level, set_remote_new_univ_level = + RemoteCounter.new_counter ~name:"Universes" 0 ~incr:((+) 1) + ~build:(fun n -> Univ.Level.make (Global.current_dirpath ()) n) + +let new_univ_level _ = new_univ_level () + (* Univ.Level.make db (new_univ_level ()) *) + +let fresh_level () = new_univ_level (Global.current_dirpath ()) + +(* TODO: remove *) +let new_univ dp = Univ.Universe.make (new_univ_level dp) +let new_Type dp = mkType (new_univ dp) +let new_Type_sort dp = Type (new_univ dp) + +let fresh_universe_instance ctx = + Instance.subst_fn (fun _ -> new_univ_level (Global.current_dirpath ())) + (UContext.instance ctx) + +let fresh_instance_from_context ctx = + let inst = fresh_universe_instance ctx in + let constraints = instantiate_univ_constraints inst ctx in + inst, constraints + +let fresh_instance ctx = + let ctx' = ref LSet.empty in + let inst = + Instance.subst_fn (fun v -> + let u = new_univ_level (Global.current_dirpath ()) in + ctx' := LSet.add u !ctx'; u) + (UContext.instance ctx) + in !ctx', inst + +let existing_instance ctx inst = + let () = + let a1 = Instance.to_array inst + and a2 = Instance.to_array (UContext.instance ctx) in + let len1 = Array.length a1 and len2 = Array.length a2 in + if not (len1 == len2) then + Errors.errorlabstrm "Universes" + (str "Polymorphic constant expected " ++ int len2 ++ + str" levels but was given " ++ int len1) + else () + in LSet.empty, inst + +let fresh_instance_from ctx inst = + let ctx', inst = + match inst with + | Some inst -> existing_instance ctx inst + | None -> fresh_instance ctx + in + let constraints = instantiate_univ_constraints inst ctx in + inst, (ctx', constraints) + +let unsafe_instance_from ctx = + (Univ.UContext.instance ctx, ctx) + +(** Fresh universe polymorphic construction *) + +let fresh_constant_instance env c inst = + let cb = lookup_constant c env in + if cb.Declarations.const_polymorphic then + let inst, ctx = + fresh_instance_from + (Declareops.universes_of_constant (Environ.opaque_tables env) cb) inst + in + ((c, inst), ctx) + else ((c,Instance.empty), ContextSet.empty) + +let fresh_inductive_instance env ind inst = + let mib, mip = Inductive.lookup_mind_specif env ind in + if mib.Declarations.mind_polymorphic then + let inst, ctx = fresh_instance_from mib.Declarations.mind_universes inst in + ((ind,inst), ctx) + else ((ind,Instance.empty), ContextSet.empty) + +let fresh_constructor_instance env (ind,i) inst = + let mib, mip = Inductive.lookup_mind_specif env ind in + if mib.Declarations.mind_polymorphic then + let inst, ctx = fresh_instance_from mib.Declarations.mind_universes inst in + (((ind,i),inst), ctx) + else (((ind,i),Instance.empty), ContextSet.empty) + +let unsafe_constant_instance env c = + let cb = lookup_constant c env in + if cb.Declarations.const_polymorphic then + let inst, ctx = unsafe_instance_from + (Declareops.universes_of_constant (Environ.opaque_tables env) cb) in + ((c, inst), ctx) + else ((c,Instance.empty), UContext.empty) + +let unsafe_inductive_instance env ind = + let mib, mip = Inductive.lookup_mind_specif env ind in + if mib.Declarations.mind_polymorphic then + let inst, ctx = unsafe_instance_from mib.Declarations.mind_universes in + ((ind,inst), ctx) + else ((ind,Instance.empty), UContext.empty) + +let unsafe_constructor_instance env (ind,i) = + let mib, mip = Inductive.lookup_mind_specif env ind in + if mib.Declarations.mind_polymorphic then + let inst, ctx = unsafe_instance_from mib.Declarations.mind_universes in + (((ind,i),inst), ctx) + else (((ind,i),Instance.empty), UContext.empty) + +open Globnames + +let fresh_global_instance ?names env gr = + match gr with + | VarRef id -> mkVar id, ContextSet.empty + | ConstRef sp -> + let c, ctx = fresh_constant_instance env sp names in + mkConstU c, ctx + | ConstructRef sp -> + let c, ctx = fresh_constructor_instance env sp names in + mkConstructU c, ctx + | IndRef sp -> + let c, ctx = fresh_inductive_instance env sp names in + mkIndU c, ctx + +let fresh_constant_instance env sp = + fresh_constant_instance env sp None + +let fresh_inductive_instance env sp = + fresh_inductive_instance env sp None + +let fresh_constructor_instance env sp = + fresh_constructor_instance env sp None + +let unsafe_global_instance env gr = + match gr with + | VarRef id -> mkVar id, UContext.empty + | ConstRef sp -> + let c, ctx = unsafe_constant_instance env sp in + mkConstU c, ctx + | ConstructRef sp -> + let c, ctx = unsafe_constructor_instance env sp in + mkConstructU c, ctx + | IndRef sp -> + let c, ctx = unsafe_inductive_instance env sp in + mkIndU c, ctx + +let constr_of_global gr = + let c, ctx = fresh_global_instance (Global.env ()) gr in + if not (Univ.ContextSet.is_empty ctx) then + if Univ.LSet.is_empty (Univ.ContextSet.levels ctx) then + (* Should be an error as we might forget constraints, allow for now + to make firstorder work with "using" clauses *) + c + else raise (Invalid_argument + ("constr_of_global: globalization of polymorphic reference " ^ + Pp.string_of_ppcmds (Nametab.pr_global_env Id.Set.empty gr) ^ + " would forget universes.")) + else c + +let constr_of_reference = constr_of_global + +let unsafe_constr_of_global gr = + unsafe_global_instance (Global.env ()) gr + +let constr_of_global_univ (gr,u) = + match gr with + | VarRef id -> mkVar id + | ConstRef sp -> mkConstU (sp,u) + | ConstructRef sp -> mkConstructU (sp,u) + | IndRef sp -> mkIndU (sp,u) + +let fresh_global_or_constr_instance env = function + | IsConstr c -> c, ContextSet.empty + | IsGlobal gr -> fresh_global_instance env gr + +let global_of_constr c = + match kind_of_term c with + | Const (c, u) -> ConstRef c, u + | Ind (i, u) -> IndRef i, u + | Construct (c, u) -> ConstructRef c, u + | Var id -> VarRef id, Instance.empty + | _ -> raise Not_found + +let global_app_of_constr c = + match kind_of_term c with + | Const (c, u) -> (ConstRef c, u), None + | Ind (i, u) -> (IndRef i, u), None + | Construct (c, u) -> (ConstructRef c, u), None + | Var id -> (VarRef id, Instance.empty), None + | Proj (p, c) -> (ConstRef (Projection.constant p), Instance.empty), Some c + | _ -> raise Not_found + +open Declarations + +let type_of_reference env r = + match r with + | VarRef id -> Environ.named_type id env, ContextSet.empty + | ConstRef c -> + let cb = Environ.lookup_constant c env in + let ty = Typeops.type_of_constant_type env cb.const_type in + if cb.const_polymorphic then + let inst, ctx = fresh_instance_from (Declareops.universes_of_constant (Environ.opaque_tables env) cb) None in + Vars.subst_instance_constr inst ty, ctx + else ty, ContextSet.empty + + | IndRef ind -> + let (mib, oib as specif) = Inductive.lookup_mind_specif env ind in + if mib.mind_polymorphic then + let inst, ctx = fresh_instance_from mib.mind_universes None in + let ty = Inductive.type_of_inductive env (specif, inst) in + ty, ctx + else + let ty = Inductive.type_of_inductive env (specif, Univ.Instance.empty) in + ty, ContextSet.empty + | ConstructRef cstr -> + let (mib,oib as specif) = Inductive.lookup_mind_specif env (inductive_of_constructor cstr) in + if mib.mind_polymorphic then + let inst, ctx = fresh_instance_from mib.mind_universes None in + Inductive.type_of_constructor (cstr,inst) specif, ctx + else Inductive.type_of_constructor (cstr,Instance.empty) specif, ContextSet.empty + +let type_of_global t = type_of_reference (Global.env ()) t + +let unsafe_type_of_reference env r = + match r with + | VarRef id -> Environ.named_type id env + | ConstRef c -> + let cb = Environ.lookup_constant c env in + Typeops.type_of_constant_type env cb.const_type + + | IndRef ind -> + let (mib, oib as specif) = Inductive.lookup_mind_specif env ind in + let (_, inst), _ = unsafe_inductive_instance env ind in + Inductive.type_of_inductive env (specif, inst) + + | ConstructRef (ind, _ as cstr) -> + let (mib,oib as specif) = Inductive.lookup_mind_specif env (inductive_of_constructor cstr) in + let (_, inst), _ = unsafe_inductive_instance env ind in + Inductive.type_of_constructor (cstr,inst) specif + +let unsafe_type_of_global t = unsafe_type_of_reference (Global.env ()) t + +let fresh_sort_in_family env = function + | InProp -> prop_sort, ContextSet.empty + | InSet -> set_sort, ContextSet.empty + | InType -> + let u = fresh_level () in + Type (Univ.Universe.make u), ContextSet.singleton u + +let new_sort_in_family sf = + fst (fresh_sort_in_family (Global.env ()) sf) + +let extend_context (a, ctx) (ctx') = + (a, ContextSet.union ctx ctx') + +let new_global_univ () = + let u = fresh_level () in + (Univ.Universe.make u, ContextSet.singleton u) + +(** Simplification *) + +module LevelUnionFind = Unionfind.Make (Univ.LSet) (Univ.LMap) + +let add_list_map u t map = + try + let l = LMap.find u map in + LMap.update u (t :: l) map + with Not_found -> + LMap.add u [t] map + +module UF = LevelUnionFind + +(** Precondition: flexible <= ctx *) +let choose_canonical ctx flexible algs s = + let global = LSet.diff s ctx in + let flexible, rigid = LSet.partition flexible (LSet.inter s ctx) in + (** If there is a global universe in the set, choose it *) + if not (LSet.is_empty global) then + let canon = LSet.choose global in + canon, (LSet.remove canon global, rigid, flexible) + else (** No global in the equivalence class, choose a rigid one *) + if not (LSet.is_empty rigid) then + let canon = LSet.choose rigid in + canon, (global, LSet.remove canon rigid, flexible) + else (** There are only flexible universes in the equivalence + class, choose a non-algebraic. *) + let algs, nonalgs = LSet.partition (fun x -> LSet.mem x algs) flexible in + if not (LSet.is_empty nonalgs) then + let canon = LSet.choose nonalgs in + canon, (global, rigid, LSet.remove canon flexible) + else + let canon = LSet.choose algs in + canon, (global, rigid, LSet.remove canon flexible) + +let subst_univs_fn_puniverses lsubst (c, u as cu) = + let u' = Instance.subst_fn lsubst u in + if u' == u then cu else (c, u') + +let nf_evars_and_universes_opt_subst f subst = + let subst = fun l -> match LMap.find l subst with None -> raise Not_found | Some l' -> l' in + let lsubst = Univ.level_subst_of subst in + let rec aux c = + match kind_of_term c with + | Evar (evk, args) -> + let args = Array.map aux args in + (match try f (evk, args) with Not_found -> None with + | None -> c + | Some c -> aux c) + | Const pu -> + let pu' = subst_univs_fn_puniverses lsubst pu in + if pu' == pu then c else mkConstU pu' + | Ind pu -> + let pu' = subst_univs_fn_puniverses lsubst pu in + if pu' == pu then c else mkIndU pu' + | Construct pu -> + let pu' = subst_univs_fn_puniverses lsubst pu in + if pu' == pu then c else mkConstructU pu' + | Sort (Type u) -> + let u' = Univ.subst_univs_universe subst u in + if u' == u then c else mkSort (sort_of_univ u') + | _ -> map_constr aux c + in aux + +let fresh_universe_context_set_instance ctx = + if ContextSet.is_empty ctx then LMap.empty, ctx + else + let (univs, cst) = ContextSet.levels ctx, ContextSet.constraints ctx in + let univs',subst = LSet.fold + (fun u (univs',subst) -> + let u' = fresh_level () in + (LSet.add u' univs', LMap.add u u' subst)) + univs (LSet.empty, LMap.empty) + in + let cst' = subst_univs_level_constraints subst cst in + subst, (univs', cst') + +let normalize_univ_variable ~find ~update = + let rec aux cur = + let b = find cur in + let b' = subst_univs_universe aux b in + if Universe.equal b' b then b + else update cur b' + in aux + +let normalize_univ_variable_opt_subst ectx = + let find l = + match Univ.LMap.find l !ectx with + | Some b -> b + | None -> raise Not_found + in + let update l b = + assert (match Universe.level b with Some l' -> not (Level.equal l l') | None -> true); + ectx := Univ.LMap.add l (Some b) !ectx; b + in normalize_univ_variable ~find ~update + +let normalize_univ_variable_subst subst = + let find l = Univ.LMap.find l !subst in + let update l b = + assert (match Universe.level b with Some l' -> not (Level.equal l l') | None -> true); + subst := Univ.LMap.add l b !subst; b in + normalize_univ_variable ~find ~update + +let normalize_universe_opt_subst subst = + let normlevel = normalize_univ_variable_opt_subst subst in + subst_univs_universe normlevel + +let normalize_universe_subst subst = + let normlevel = normalize_univ_variable_subst subst in + subst_univs_universe normlevel + +let normalize_opt_subst ctx = + let ectx = ref ctx in + let normalize = normalize_univ_variable_opt_subst ectx in + let () = + Univ.LMap.iter (fun u v -> + if Option.is_empty v then () + else try ignore(normalize u) with Not_found -> assert(false)) ctx + in !ectx + +type universe_opt_subst = universe option universe_map + +let make_opt_subst s = + fun x -> + (match Univ.LMap.find x s with + | Some u -> u + | None -> raise Not_found) + +let subst_opt_univs_constr s = + let f = make_opt_subst s in + Vars.subst_univs_fn_constr f + + +let normalize_univ_variables ctx = + let ctx = normalize_opt_subst ctx in + let undef, def, subst = + Univ.LMap.fold (fun u v (undef, def, subst) -> + match v with + | None -> (Univ.LSet.add u undef, def, subst) + | Some b -> (undef, Univ.LSet.add u def, Univ.LMap.add u b subst)) + ctx (Univ.LSet.empty, Univ.LSet.empty, Univ.LMap.empty) + in ctx, undef, def, subst + +let pr_universe_body = function + | None -> mt () + | Some v -> str" := " ++ Univ.Universe.pr v + +let pr_universe_opt_subst = Univ.LMap.pr pr_universe_body + +exception Found of Level.t +let find_inst insts v = + try LMap.iter (fun k (enf,alg,v') -> + if not alg && enf && Universe.equal v' v then raise (Found k)) + insts; raise Not_found + with Found l -> l + +let compute_lbound left = + (** The universe variable was not fixed yet. + Compute its level using its lower bound. *) + let sup l lbound = + match lbound with + | None -> Some l + | Some l' -> Some (Universe.sup l l') + in + List.fold_left (fun lbound (d, l) -> + if d == Le (* l <= ?u *) then sup l lbound + else (* l < ?u *) + (assert (d == Lt); + if not (Universe.level l == None) then + sup (Universe.super l) lbound + else None)) + None left + +let instantiate_with_lbound u lbound alg enforce (ctx, us, algs, insts, cstrs) = + if enforce then + let inst = Universe.make u in + let cstrs' = enforce_leq lbound inst cstrs in + (ctx, us, LSet.remove u algs, + LMap.add u (enforce,alg,lbound) insts, cstrs'), (enforce, alg, inst) + else (* Actually instantiate *) + (Univ.LSet.remove u ctx, Univ.LMap.add u (Some lbound) us, algs, + LMap.add u (enforce,alg,lbound) insts, cstrs), (enforce, alg, lbound) + +type constraints_map = (Univ.constraint_type * Univ.LMap.key) list Univ.LMap.t + +let pr_constraints_map cmap = + LMap.fold (fun l cstrs acc -> + Level.pr l ++ str " => " ++ + prlist_with_sep spc (fun (d,r) -> pr_constraint_type d ++ Level.pr r) cstrs ++ fnl () + ++ acc) + cmap (mt ()) + +let minimize_univ_variables ctx us algs left right cstrs = + let left, lbounds = + Univ.LMap.fold (fun r lower (left, lbounds as acc) -> + if Univ.LMap.mem r us || not (Univ.LSet.mem r ctx) then acc + else (* Fixed universe, just compute its glb for sharing *) + let lbounds' = + match compute_lbound (List.map (fun (d,l) -> d, Universe.make l) lower) with + | None -> lbounds + | Some lbound -> LMap.add r (true, false, lbound) lbounds + in (Univ.LMap.remove r left, lbounds')) + left (left, Univ.LMap.empty) + in + let rec instance (ctx', us, algs, insts, cstrs as acc) u = + let acc, left = + try let l = LMap.find u left in + List.fold_left (fun (acc, left') (d, l) -> + let acc', (enf,alg,l') = aux acc l in + (* if alg then assert(not alg); *) + let l' = + if enf then Universe.make l + else l' + (* match Universe.level l' with Some _ -> l' | None -> Universe.make l *) + in + acc', (d, l') :: left') (acc, []) l + with Not_found -> acc, [] + and right = + try Some (LMap.find u right) + with Not_found -> None + in + let instantiate_lbound lbound = + let alg = LSet.mem u algs in + if alg then + (* u is algebraic and has no upper bound constraints: we + instantiate it with it's lower bound, if any *) + instantiate_with_lbound u lbound true false acc + else (* u is non algebraic *) + match Universe.level lbound with + | Some l -> (* The lowerbound is directly a level *) + (* u is not algebraic but has no upper bounds, + we instantiate it with its lower bound if it is a + different level, otherwise we keep it. *) + if not (Level.equal l u) && not (LSet.mem l algs) then + (* if right = None then. Should check that u does not + have upper constraints that are not already in right *) + instantiate_with_lbound u lbound false false acc + (* else instantiate_with_lbound u lbound false true acc *) + else + (* assert false: l can't be alg *) + acc, (true, false, lbound) + | None -> + try + (* if right <> None then raise Not_found; *) + (* Another universe represents the same lower bound, + we can share them with no harm. *) + let can = find_inst insts lbound in + instantiate_with_lbound u (Universe.make can) false false acc + with Not_found -> + (* We set u as the canonical universe representing lbound *) + instantiate_with_lbound u lbound false true acc + in + let acc' acc = + match right with + | None -> acc + | Some cstrs -> + let dangling = List.filter (fun (d, r) -> not (LMap.mem r us)) cstrs in + if List.is_empty dangling then acc + else + let ((ctx', us, algs, insts, cstrs), (enf,_,inst as b)) = acc in + let cstrs' = List.fold_left (fun cstrs (d, r) -> + if d == Univ.Le then + enforce_leq inst (Universe.make r) cstrs + else + try let lev = Option.get (Universe.level inst) in + Constraint.add (lev, d, r) cstrs + with Option.IsNone -> assert false) + cstrs dangling + in + (ctx', us, algs, insts, cstrs'), b + in + if not (LSet.mem u ctx) then acc' (acc, (true, false, Universe.make u)) + else + let lbound = compute_lbound left in + match lbound with + | None -> (* Nothing to do *) + acc' (acc, (true, false, Universe.make u)) + | Some lbound -> + acc' (instantiate_lbound lbound) + and aux (ctx', us, algs, seen, cstrs as acc) u = + try acc, LMap.find u seen + with Not_found -> instance acc u + in + LMap.fold (fun u v (ctx', us, algs, seen, cstrs as acc) -> + if v == None then fst (aux acc u) + else LSet.remove u ctx', us, LSet.remove u algs, seen, cstrs) + us (ctx, us, algs, lbounds, cstrs) + +let normalize_context_set ctx us algs = + let (ctx, csts) = ContextSet.levels ctx, ContextSet.constraints ctx in + let uf = UF.create () in + let csts = + (* We first put constraints in a normal-form: all self-loops are collapsed + to equalities. *) + let g = Univ.merge_constraints csts Univ.empty_universes in + Univ.constraints_of_universes g + in + let noneqs = + Constraint.fold (fun (l,d,r) noneqs -> + if d == Eq then (UF.union l r uf; noneqs) + else Constraint.add (l,d,r) noneqs) + csts Constraint.empty + in + let partition = UF.partition uf in + let flex x = LMap.mem x us in + let ctx, subst, eqs = List.fold_left (fun (ctx, subst, cstrs) s -> + let canon, (global, rigid, flexible) = choose_canonical ctx flex algs s in + (* Add equalities for globals which can't be merged anymore. *) + let cstrs = LSet.fold (fun g cst -> + Constraint.add (canon, Univ.Eq, g) cst) global + cstrs + in + let subst = LSet.fold (fun f -> LMap.add f canon) rigid subst in + let subst = LSet.fold (fun f -> LMap.add f canon) flexible subst in + (LSet.diff (LSet.diff ctx rigid) flexible, subst, cstrs)) + (ctx, LMap.empty, Constraint.empty) partition + in + (* Noneqs is now in canonical form w.r.t. equality constraints, + and contains only inequality constraints. *) + let noneqs = subst_univs_level_constraints subst noneqs in + let us = LMap.fold (fun u v acc -> LMap.add u (Some (Universe.make v)) acc) subst us in + (* Compute the left and right set of flexible variables, constraints + mentionning other variables remain in noneqs. *) + let noneqs, ucstrsl, ucstrsr = + Constraint.fold (fun (l,d,r as cstr) (noneq, ucstrsl, ucstrsr) -> + let lus = LMap.mem l us + and rus = LMap.mem r us + in + let ucstrsl' = + if lus then add_list_map l (d, r) ucstrsl + else ucstrsl + and ucstrsr' = + add_list_map r (d, l) ucstrsr + in + let noneqs = + if lus || rus then noneq + else Constraint.add cstr noneq + in (noneqs, ucstrsl', ucstrsr')) + noneqs (Constraint.empty, LMap.empty, LMap.empty) + in + (* Now we construct the instantiation of each variable. *) + let ctx', us, algs, inst, noneqs = + minimize_univ_variables ctx us algs ucstrsr ucstrsl noneqs + in + let us = normalize_opt_subst us in + (us, algs), (ctx', Constraint.union noneqs eqs) + +(* let normalize_conkey = Profile.declare_profile "normalize_context_set" *) +(* let normalize_context_set a b c = Profile.profile3 normalize_conkey normalize_context_set a b c *) + +let universes_of_constr c = + let rec aux s c = + match kind_of_term c with + | Const (_, u) | Ind (_, u) | Construct (_, u) -> + LSet.union (Instance.levels u) s + | Sort u when not (Sorts.is_small u) -> + let u = univ_of_sort u in + LSet.union (Universe.levels u) s + | _ -> fold_constr aux s c + in aux LSet.empty c + +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 + +let simplify_universe_context (univs,csts) = + let uf = UF.create () in + let noneqs = + Constraint.fold (fun (l,d,r) noneqs -> + if d == Eq && (LSet.mem l univs || LSet.mem r univs) then + (UF.union l r uf; noneqs) + else Constraint.add (l,d,r) noneqs) + csts Constraint.empty + in + let partition = UF.partition uf in + let flex x = LSet.mem x univs in + let subst, univs', csts' = List.fold_left (fun (subst, univs, cstrs) s -> + let canon, (global, rigid, flexible) = choose_canonical univs flex LSet.empty s in + (* Add equalities for globals which can't be merged anymore. *) + let cstrs = LSet.fold (fun g cst -> + Constraint.add (canon, Univ.Eq, g) cst) (LSet.union global rigid) + cstrs + in + let subst = LSet.fold (fun f -> LMap.add f canon) + flexible subst + in (subst, LSet.diff univs flexible, cstrs)) + (LMap.empty, univs, noneqs) partition + in + (* Noneqs is now in canonical form w.r.t. equality constraints, + and contains only inequality constraints. *) + let csts' = subst_univs_level_constraints subst csts' in + (univs', csts'), subst + +let is_small_leq (l,d,r) = + Level.is_small l && d == Univ.Le + +(* Prop < i <-> Set+1 <= i <-> Set < i *) +let translate_cstr (l,d,r as cstr) = + if Level.equal Level.prop l && d == Univ.Lt then + (Level.set, d, r) + else cstr + +let refresh_constraints univs (ctx, cstrs) = + let cstrs', univs' = + Univ.Constraint.fold (fun c (cstrs', univs as acc) -> + let c = translate_cstr c in + if Univ.check_constraint univs c && not (is_small_leq c) then acc + else (Univ.Constraint.add c cstrs', Univ.enforce_constraint c univs)) + cstrs (Univ.Constraint.empty, univs) + in ((ctx, cstrs'), univs') + + +(**********************************************************************) +(* Tools for sort-polymorphic inductive types *) + +(* Miscellaneous functions to remove or test local univ assumed to + occur only in the le constraints *) + +(* + Solve a system of universe constraint of the form + + u_s11, ..., u_s1p1, w1 <= u1 + ... + u_sn1, ..., u_snpn, wn <= un + +where + + - the ui (1 <= i <= n) are universe variables, + - the sjk select subsets of the ui for each equations, + - the wi are arbitrary complex universes that do not mention the ui. +*) + +let is_direct_sort_constraint s v = match s with + | Some u -> univ_level_mem u v + | None -> false + +let solve_constraints_system levels level_bounds level_min = + let open Univ in + let levels = + Array.mapi (fun i o -> + match o with + | Some u -> + (match Universe.level u with + | Some u -> Some u + | _ -> level_bounds.(i) <- Universe.sup level_bounds.(i) u; None) + | None -> None) + levels in + let v = Array.copy level_bounds in + let nind = Array.length v in + let clos = Array.map (fun _ -> Int.Set.empty) levels in + (* First compute the transitive closure of the levels dependencies *) + for i=0 to nind-1 do + for j=0 to nind-1 do + if not (Int.equal i j) && is_direct_sort_constraint levels.(j) v.(i) then + clos.(i) <- Int.Set.add j clos.(i); + done; + done; + let rec closure () = + let continue = ref false in + Array.iteri (fun i deps -> + let deps' = + Int.Set.fold (fun j acc -> Int.Set.union acc clos.(j)) deps deps + in + if Int.Set.equal deps deps' then () + else (clos.(i) <- deps'; continue := true)) + clos; + if !continue then closure () + else () + in + closure (); + for i=0 to nind-1 do + for j=0 to nind-1 do + if not (Int.equal i j) && Int.Set.mem j clos.(i) then + (v.(i) <- Universe.sup v.(i) level_bounds.(j); + level_min.(i) <- Universe.sup level_min.(i) level_min.(j)) + done; + for j=0 to nind-1 do + match levels.(j) with + | Some u -> v.(i) <- univ_level_rem u v.(i) level_min.(i) + | None -> () + done + done; + v diff --git a/library/universes.mli b/library/universes.mli new file mode 100644 index 00000000..f2f68d32 --- /dev/null +++ b/library/universes.mli @@ -0,0 +1,253 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* universe_names +val set_global_universe_names : universe_names -> unit + +val pr_with_global_universes : Level.t -> Pp.std_ppcmds + +(** The global universe counter *) +val set_remote_new_univ_level : universe_level RemoteCounter.installer + +(** Side-effecting functions creating new universe levels. *) + +val new_univ_level : Names.dir_path -> universe_level +val new_univ : Names.dir_path -> universe +val new_Type : Names.dir_path -> types +val new_Type_sort : Names.dir_path -> sorts + +val new_global_univ : unit -> universe in_universe_context_set +val new_sort_in_family : sorts_family -> sorts + +(** {6 Constraints for type inference} + + When doing conversion of universes, not only do we have =/<= constraints but + also Lub constraints which correspond to unification of two levels which might + not be necessary if unfolding is performed. +*) + +type universe_constraint_type = ULe | UEq | ULub + +type universe_constraint = universe * universe_constraint_type * universe +module Constraints : sig + include Set.S with type elt = universe_constraint + + val pr : t -> Pp.std_ppcmds +end + +type universe_constraints = Constraints.t +type 'a universe_constrained = 'a * universe_constraints +type 'a universe_constraint_function = 'a -> 'a -> universe_constraints -> universe_constraints + +val subst_univs_universe_constraints : universe_subst_fn -> + universe_constraints -> universe_constraints + +val enforce_eq_instances_univs : bool -> universe_instance universe_constraint_function + +val to_constraints : universes -> universe_constraints -> constraints + +(** [eq_constr_univs_infer u a b] is [true, c] if [a] equals [b] modulo alpha, casts, + application grouping, the universe constraints in [u] and additional constraints [c]. *) +val eq_constr_univs_infer : Univ.universes -> constr -> constr -> bool universe_constrained + +(** [leq_constr_univs u a b] is [true, c] if [a] is convertible to [b] + modulo alpha, casts, application grouping, the universe constraints + in [u] and additional constraints [c]. *) +val leq_constr_univs_infer : Univ.universes -> constr -> constr -> bool universe_constrained + +(** [eq_constr_universes a b] [true, c] if [a] equals [b] modulo alpha, casts, + application grouping and the universe constraints in [c]. *) +val eq_constr_universes : constr -> constr -> bool universe_constrained + +(** [leq_constr_universes a b] [true, c] if [a] is convertible to [b] modulo + alpha, casts, application grouping and the universe constraints in [c]. *) +val leq_constr_universes : constr -> constr -> bool universe_constrained + +(** [eq_constr_universes a b] [true, c] if [a] equals [b] modulo alpha, casts, + application grouping and the universe constraints in [c]. *) +val eq_constr_universes_proj : env -> constr -> constr -> bool universe_constrained + +(** Build a fresh instance for a given context, its associated substitution and + the instantiated constraints. *) + +val fresh_instance_from_context : universe_context -> + universe_instance constrained + +val fresh_instance_from : universe_context -> universe_instance option -> + universe_instance in_universe_context_set + +val fresh_sort_in_family : env -> sorts_family -> + sorts in_universe_context_set +val fresh_constant_instance : env -> constant -> + pconstant in_universe_context_set +val fresh_inductive_instance : env -> inductive -> + pinductive in_universe_context_set +val fresh_constructor_instance : env -> constructor -> + pconstructor in_universe_context_set + +val fresh_global_instance : ?names:Univ.Instance.t -> env -> Globnames.global_reference -> + constr in_universe_context_set + +val fresh_global_or_constr_instance : env -> Globnames.global_reference_or_constr -> + constr in_universe_context_set + +(** Get fresh variables for the universe context. + Useful to make tactics that manipulate constrs in universe contexts polymorphic. *) +val fresh_universe_context_set_instance : universe_context_set -> + universe_level_subst * universe_context_set + +(** Raises [Not_found] if not a global reference. *) +val global_of_constr : constr -> Globnames.global_reference puniverses + +val global_app_of_constr : constr -> Globnames.global_reference puniverses * constr option + +val constr_of_global_univ : Globnames.global_reference puniverses -> constr + +val extend_context : 'a in_universe_context_set -> universe_context_set -> + 'a in_universe_context_set + +(** Simplification and pruning of constraints: + [normalize_context_set ctx us] + + - Instantiate the variables in [us] with their most precise + universe levels respecting the constraints. + + - Normalizes the context [ctx] w.r.t. equality constraints, + choosing a canonical universe in each equivalence class + (a global one if there is one) and transitively saturate + the constraints w.r.t to the equalities. *) + +module UF : Unionfind.PartitionSig with type elt = universe_level + +type universe_opt_subst = universe option universe_map + +val make_opt_subst : universe_opt_subst -> universe_subst_fn + +val subst_opt_univs_constr : universe_opt_subst -> constr -> constr + +val normalize_context_set : universe_context_set -> + universe_opt_subst (* The defined and undefined variables *) -> + universe_set (* univ variables that can be substituted by algebraics *) -> + (universe_opt_subst * universe_set) in_universe_context_set + +val normalize_univ_variables : universe_opt_subst -> + universe_opt_subst * universe_set * universe_set * universe_subst + +val normalize_univ_variable : + find:(universe_level -> universe) -> + update:(universe_level -> universe -> universe) -> + universe_level -> universe + +val normalize_univ_variable_opt_subst : universe_opt_subst ref -> + (universe_level -> universe) + +val normalize_univ_variable_subst : universe_subst ref -> + (universe_level -> universe) + +val normalize_universe_opt_subst : universe_opt_subst ref -> + (universe -> universe) + +val normalize_universe_subst : universe_subst ref -> + (universe -> universe) + +(** Create a fresh global in the global environment, without side effects. + BEWARE: this raises an ANOMALY on polymorphic constants/inductives: + the constraints should be properly added to an evd. + See Evd.fresh_global, Evarutil.new_global, and pf_constr_of_global for + the proper way to get a fresh copy of a global reference. *) +val constr_of_global : Globnames.global_reference -> constr + +(** ** DEPRECATED ** synonym of [constr_of_global] *) +val constr_of_reference : Globnames.global_reference -> constr + +(** [unsafe_constr_of_global gr] turns [gr] into a constr, works on polymorphic + references by taking the original universe instance that is not recorded + anywhere. The constraints are forgotten as well. DO NOT USE in new code. *) +val unsafe_constr_of_global : Globnames.global_reference -> constr in_universe_context + +(** Returns the type of the global reference, by creating a fresh instance of polymorphic + references and computing their instantiated universe context. (side-effect on the + universe counter, use with care). *) +val type_of_global : Globnames.global_reference -> types in_universe_context_set + +(** [unsafe_type_of_global gr] returns [gr]'s type, works on polymorphic + references by taking the original universe instance that is not recorded + anywhere. The constraints are forgotten as well. + USE with care. *) +val unsafe_type_of_global : Globnames.global_reference -> types + +(** Full universes substitutions into terms *) + +val nf_evars_and_universes_opt_subst : (existential -> constr option) -> + universe_opt_subst -> constr -> constr + +(** Shrink a universe context to a restricted set of variables *) + +val universes_of_constr : constr -> universe_set +val restrict_universe_context : universe_context_set -> universe_set -> universe_context_set +val simplify_universe_context : universe_context_set -> + universe_context_set * universe_level_subst + +val refresh_constraints : universes -> universe_context_set -> universe_context_set * universes + +(** Pretty-printing *) + +val pr_universe_opt_subst : universe_opt_subst -> Pp.std_ppcmds + +(* For tracing *) + +type constraints_map = (Univ.constraint_type * Univ.LMap.key) list Univ.LMap.t + +val pr_constraints_map : constraints_map -> Pp.std_ppcmds + +val choose_canonical : universe_set -> (Level.t -> bool) (* flexibles *) -> universe_set -> universe_set -> + universe_level * (universe_set * universe_set * universe_set) + +val compute_lbound : (constraint_type * Univ.universe) list -> universe option + +val instantiate_with_lbound : + Univ.LMap.key -> + Univ.universe -> + bool -> + bool -> + Univ.LSet.t * Univ.universe option Univ.LMap.t * + Univ.LSet.t * + (bool * bool * Univ.universe) Univ.LMap.t * Univ.constraints -> + (Univ.LSet.t * Univ.universe option Univ.LMap.t * + Univ.LSet.t * + (bool * bool * Univ.universe) Univ.LMap.t * Univ.constraints) * + (bool * bool * Univ.universe) + +val minimize_univ_variables : + Univ.LSet.t -> + Univ.universe option Univ.LMap.t -> + Univ.LSet.t -> + constraints_map -> constraints_map -> + Univ.constraints -> + Univ.LSet.t * Univ.universe option Univ.LMap.t * + Univ.LSet.t * + (bool * bool * Univ.universe) Univ.LMap.t * Univ.constraints + +(** {6 Support for old-style sort-polymorphism } *) + +val solve_constraints_system : universe option array -> universe array -> universe array -> + universe array -- cgit v1.2.3