diff options
Diffstat (limited to 'vernac')
-rw-r--r-- | vernac/classes.ml | 11 | ||||
-rw-r--r-- | vernac/command.ml | 72 | ||||
-rw-r--r-- | vernac/command.mli | 6 | ||||
-rw-r--r-- | vernac/declareDef.ml | 3 | ||||
-rw-r--r-- | vernac/lemmas.ml | 35 | ||||
-rw-r--r-- | vernac/lemmas.mli | 2 | ||||
-rw-r--r-- | vernac/locality.ml | 57 | ||||
-rw-r--r-- | vernac/locality.mli | 13 | ||||
-rw-r--r-- | vernac/obligations.ml | 7 | ||||
-rw-r--r-- | vernac/record.ml | 5 | ||||
-rw-r--r-- | vernac/record.mli | 2 | ||||
-rw-r--r-- | vernac/vernacentries.ml | 386 | ||||
-rw-r--r-- | vernac/vernacentries.mli | 2 | ||||
-rw-r--r-- | vernac/vernacinterp.ml | 15 | ||||
-rw-r--r-- | vernac/vernacinterp.mli | 16 | ||||
-rw-r--r-- | vernac/vernacstate.ml | 30 | ||||
-rw-r--r-- | vernac/vernacstate.mli | 2 |
17 files changed, 348 insertions, 316 deletions
diff --git a/vernac/classes.ml b/vernac/classes.ml index b80741269..3e47f881c 100644 --- a/vernac/classes.ml +++ b/vernac/classes.ml @@ -114,9 +114,10 @@ let instance_hook k info global imps ?hook cst = let declare_instance_constant k info global imps ?hook id decl poly evm term termtype = let kind = IsDefinition Instance in - let evm = - let levels = Univ.LSet.union (Univops.universes_of_constr termtype) - (Univops.universes_of_constr term) in + let evm = + let env = Global.env () in + let levels = Univ.LSet.union (Univops.universes_of_constr env termtype) + (Univops.universes_of_constr env term) in Evd.restrict_universe_context evm levels in let uctx = Evd.check_univ_decl ~poly evm decl in @@ -126,7 +127,7 @@ let declare_instance_constant k info global imps ?hook id decl poly evm term ter let cdecl = (DefinitionEntry entry, kind) in let kn = Declare.declare_constant id cdecl in Declare.definition_message id; - Universes.register_universe_binders (ConstRef kn) (Evd.universe_binders evm); + Declare.declare_univ_binders (ConstRef kn) (Evd.universe_binders evm); instance_hook k info global imps ?hook (ConstRef kn); id @@ -208,7 +209,7 @@ let new_instance ?(abstract=false) ?(global=false) ?(refine= !refine_instance) (ParameterEntry (None,(termtype,univs),None), Decl_kinds.IsAssumption Decl_kinds.Logical) in - Universes.register_universe_binders (ConstRef cst) (Evd.universe_binders !evars); + Declare.declare_univ_binders (ConstRef cst) (Evd.universe_binders !evars); instance_hook k pri global imps ?hook (ConstRef cst); id end else ( diff --git a/vernac/command.ml b/vernac/command.ml index 373e5a1be..23be2c308 100644 --- a/vernac/command.ml +++ b/vernac/command.ml @@ -95,7 +95,7 @@ let interp_definition pl bl poly red_option c ctypopt = let impls, ((env_bl, ctx), imps1) = interp_context_evars env evdref bl in let ctx = List.map (fun d -> map_rel_decl EConstr.Unsafe.to_constr d) ctx in let nb_args = Context.Rel.nhyps ctx in - let imps,pl,ce = + let imps,ce = match ctypopt with None -> let subst = evd_comb0 Evd.nf_univ_variables evdref in @@ -105,11 +105,10 @@ let interp_definition pl bl poly red_option c ctypopt = let c = EConstr.Unsafe.to_constr c in let nf,subst = Evarutil.e_nf_evars_and_universes evdref in let body = nf (it_mkLambda_or_LetIn c ctx) in - let vars = Univops.universes_of_constr body in - let evd = Evd.restrict_universe_context !evdref vars in - let uctx = Evd.check_univ_decl ~poly evd decl in - let binders = Evd.universe_binders evd in - imps1@(Impargs.lift_implicits nb_args imps2), binders, + let vars = EConstr.universes_of_constr env !evdref (EConstr.of_constr body) in + let () = evdref := Evd.restrict_universe_context !evdref vars in + let uctx = Evd.check_univ_decl ~poly !evdref decl in + imps1@(Impargs.lift_implicits nb_args imps2), definition_entry ~univs:uctx body | Some ctyp -> let ty, impsty = interp_type_evars_impls ~impls env_bl evdref ctyp in @@ -131,23 +130,22 @@ let interp_definition pl bl poly red_option c ctypopt = in if not (try List.for_all chk imps2 with Not_found -> false) then warn_implicits_in_term (); - let vars = Univ.LSet.union (Univops.universes_of_constr body) - (Univops.universes_of_constr typ) in - let ctx = Evd.restrict_universe_context !evdref vars in - let uctx = Evd.check_univ_decl ~poly ctx decl in - let binders = Evd.universe_binders evd in - imps1@(Impargs.lift_implicits nb_args impsty), binders, - definition_entry ~types:typ - ~univs:uctx body + let bodyvars = EConstr.universes_of_constr env !evdref (EConstr.of_constr body) in + let tyvars = EConstr.universes_of_constr env !evdref (EConstr.of_constr ty) in + let vars = Univ.LSet.union bodyvars tyvars in + let () = evdref := Evd.restrict_universe_context !evdref vars in + let uctx = Evd.check_univ_decl ~poly !evdref decl in + imps1@(Impargs.lift_implicits nb_args impsty), + definition_entry ~types:typ ~univs:uctx body in - red_constant_entry (Context.Rel.length ctx) ce !evdref red_option, !evdref, decl, pl, imps + (red_constant_entry (Context.Rel.length ctx) ce !evdref red_option, !evdref, decl, imps) -let check_definition (ce, evd, _, _, imps) = +let check_definition (ce, evd, _, imps) = check_evars_are_solved (Global.env ()) evd Evd.empty; ce let do_definition ident k univdecl bl red_option c ctypopt hook = - let (ce, evd, univdecl, pl', imps as def) = + let (ce, evd, univdecl, imps as def) = interp_definition univdecl bl (pi2 k) red_option c ctypopt in if Flags.is_program_mode () then @@ -168,12 +166,30 @@ let do_definition ident k univdecl bl red_option c ctypopt hook = ignore(Obligations.add_definition ident ~term:c cty ctx ~univdecl ~implicits:imps ~kind:k ~hook obls) else let ce = check_definition def in - ignore(DeclareDef.declare_definition ident k ce pl' imps + ignore(DeclareDef.declare_definition ident k ce (Evd.universe_binders evd) imps (Lemmas.mk_hook (fun l r -> Lemmas.call_hook (fun exn -> exn) hook l r;r))) (* 2| Variable/Hypothesis/Parameter/Axiom declarations *) +let axiom_into_instance = ref false + +let _ = + let open Goptions in + declare_bool_option + { optdepr = false; + optname = "automatically declare axioms whose type is a typeclass as instances"; + optkey = ["Typeclasses";"Axioms";"Are";"Instances"]; + optread = (fun _ -> !axiom_into_instance); + optwrite = (:=) axiom_into_instance; } + +let should_axiom_into_instance = function + | Discharge -> + (* The typeclass behaviour of Variable and Context doesn't depend + on section status *) + true + | Global | Local -> !axiom_into_instance + let declare_assumption is_coe (local,p,kind) (c,ctx) pl imps impl nl (_,ident) = match local with | Discharge when Lib.sections_are_opened () -> @@ -195,6 +211,7 @@ match local with (r,Univ.Instance.empty,true) | Global | Local | Discharge -> + let do_instance = should_axiom_into_instance local in let local = DeclareDef.get_locality ident ~kind:"axiom" local in let inl = match nl with | NoInline -> None @@ -205,9 +222,9 @@ match local with let kn = declare_constant ident ~local decl in let gr = ConstRef kn in let () = maybe_declare_manual_implicits false gr imps in - let () = Universes.register_universe_binders gr pl in + let () = Declare.declare_univ_binders gr pl in let () = assumption_message ident in - let () = Typeclasses.declare_instance None false gr in + let () = if do_instance then Typeclasses.declare_instance None false gr in let () = if is_coe then Class.try_add_new_coercion gr ~local p in let inst = match ctx with | Polymorphic_const_entry ctx -> Univ.UContext.instance ctx @@ -298,7 +315,7 @@ let do_assumptions kind nl l = let nf_evar c = EConstr.to_constr evd (EConstr.of_constr c) in let uvars, l = List.fold_left_map (fun uvars (coe,t,imps) -> let t = nf_evar t in - let uvars = Univ.LSet.union uvars (Univops.universes_of_constr t) in + let uvars = Univ.LSet.union uvars (Univops.universes_of_constr env t) in uvars, (coe,t,imps)) Univ.LSet.empty l in @@ -693,7 +710,7 @@ let declare_mutual_inductive_with_eliminations mie pl impls = let ind = (mind,i) in let gr = IndRef ind in maybe_declare_manual_implicits false gr indimpls; - Universes.register_universe_binders gr pl; + Declare.declare_univ_binders gr pl; List.iteri (fun j impls -> maybe_declare_manual_implicits false @@ -1171,7 +1188,7 @@ let declare_fixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) ind let env = Global.env() in let indexes = search_guard env indexes fixdecls in let fiximps = List.map (fun (n,r,p) -> r) fiximps in - let vars = Univops.universes_of_constr (mkFix ((indexes,0),fixdecls)) in + let vars = Univops.universes_of_constr env (mkFix ((indexes,0),fixdecls)) in let fixdecls = List.map_i (fun i _ -> mkFix ((indexes,i),fixdecls)) 0 fixnames in let evd = Evd.from_ctx ctx in @@ -1204,7 +1221,8 @@ let declare_cofixpoint local poly ((fixnames,fixdefs,fixtypes),pl,ctx,fiximps) n let fixdefs = List.map Option.get fixdefs in let fixdecls = prepare_recursive_declaration fixnames fixtypes fixdefs in let fixdecls = List.map_i (fun i _ -> mkCoFix (i,fixdecls)) 0 fixnames in - let vars = Univops.universes_of_constr (List.hd fixdecls) in + let env = Global.env () in + let vars = Univops.universes_of_constr env (List.hd fixdecls) in let fixdecls = List.map Safe_typing.mk_pure_proof fixdecls in let fiximps = List.map (fun (len,imps,idx) -> imps) fiximps in let evd = Evd.from_ctx ctx in @@ -1249,7 +1267,7 @@ let collect_evars_of_term evd c ty = Evar.Set.fold (fun ev acc -> Evd.add acc ev (Evd.find_undefined evd ev)) evars (Evd.from_ctx (Evd.evar_universe_context evd)) -let do_program_recursive local p fixkind fixl ntns = +let do_program_recursive local poly fixkind fixl ntns = let isfix = fixkind != Obligations.IsCoFixpoint in let (env, rec_sign, pl, evd), fix, info = interp_recursive isfix fixl ntns @@ -1291,8 +1309,8 @@ let do_program_recursive local p fixkind fixl ntns = end in let ctx = Evd.evar_universe_context evd in let kind = match fixkind with - | Obligations.IsFixpoint _ -> (local, p, Fixpoint) - | Obligations.IsCoFixpoint -> (local, p, CoFixpoint) + | Obligations.IsFixpoint _ -> (local, poly, Fixpoint) + | Obligations.IsCoFixpoint -> (local, poly, CoFixpoint) in Obligations.add_mutual_definitions defs ~kind ~univdecl:pl ctx ntns fixkind diff --git a/vernac/command.mli b/vernac/command.mli index 070f3e112..c7342e6da 100644 --- a/vernac/command.mli +++ b/vernac/command.mli @@ -28,7 +28,7 @@ val do_constraint : polymorphic -> val interp_definition : Vernacexpr.universe_decl_expr option -> local_binder_expr list -> polymorphic -> red_expr option -> constr_expr -> constr_expr option -> Safe_typing.private_constants definition_entry * Evd.evar_map * - Univdecls.universe_decl * Universes.universe_binders * Impargs.manual_implicits + Univdecls.universe_decl * Impargs.manual_implicits val do_definition : Id.t -> definition_kind -> Vernacexpr.universe_decl_expr option -> local_binder_expr list -> red_expr option -> constr_expr -> @@ -82,7 +82,7 @@ type one_inductive_impls = val interp_mutual_inductive : structured_inductive_expr -> decl_notation list -> cumulative_inductive_flag -> - polymorphic -> private_flag -> Decl_kinds.recursivity_kind -> + polymorphic -> private_flag -> Declarations.recursivity_kind -> mutual_inductive_entry * Universes.universe_binders * one_inductive_impls list (** Registering a mutual inductive definition together with its @@ -96,7 +96,7 @@ val declare_mutual_inductive_with_eliminations : val do_mutual_inductive : (one_inductive_expr * decl_notation list) list -> cumulative_inductive_flag -> - polymorphic -> private_flag -> Decl_kinds.recursivity_kind -> unit + polymorphic -> private_flag -> Declarations.recursivity_kind -> unit (** {6 Fixpoints and cofixpoints} *) diff --git a/vernac/declareDef.ml b/vernac/declareDef.ml index 980db4109..dfac78c04 100644 --- a/vernac/declareDef.ml +++ b/vernac/declareDef.ml @@ -36,7 +36,7 @@ let declare_global_definition ident ce local k pl imps = let kn = declare_constant ident ~local (DefinitionEntry ce, IsDefinition k) in let gr = ConstRef kn in let () = maybe_declare_manual_implicits false gr imps in - let () = Universes.register_universe_binders gr pl in + let () = Declare.declare_univ_binders gr pl in let () = definition_message ident in gr @@ -49,6 +49,7 @@ let declare_definition ident (local, p, k) ce pl imps hook = let () = definition_message ident in let gr = VarRef ident in let () = maybe_declare_manual_implicits false gr imps in + let () = Declare.declare_univ_binders gr pl in let () = if Proof_global.there_are_pending_proofs () then warn_definition_not_visible ident in diff --git a/vernac/lemmas.ml b/vernac/lemmas.ml index 42631a15b..200c2260e 100644 --- a/vernac/lemmas.ml +++ b/vernac/lemmas.ml @@ -177,7 +177,7 @@ let look_for_possibly_mutual_statements = function (* Saving a goal *) -let save ?export_seff id const cstrs pl do_guard (locality,poly,kind) hook = +let save ?export_seff id const uctx do_guard (locality,poly,kind) hook = let fix_exn = Future.fix_exn_of const.Entries.const_entry_body in try let const = adjust_guardness_conditions const do_guard in @@ -204,7 +204,7 @@ let save ?export_seff id const cstrs pl do_guard (locality,poly,kind) hook = (locality, ConstRef kn) in definition_message id; - Universes.register_universe_binders r (Option.default Universes.empty_binders pl); + Declare.declare_univ_binders r (UState.universe_binders uctx); call_hook (fun exn -> exn) hook l r with e when CErrors.noncritical e -> let e = CErrors.push e in @@ -286,17 +286,17 @@ let save_hook = ref ignore let set_save_hook f = save_hook := f let save_named ?export_seff proof = - let id,const,(cstrs,pl),do_guard,persistence,hook = proof in - save ?export_seff id const cstrs pl do_guard persistence hook + let id,const,uctx,do_guard,persistence,hook = proof in + save ?export_seff id const uctx do_guard persistence hook let check_anonymity id save_ident = if not (String.equal (atompart_of_id id) (Id.to_string (default_thm_id))) then user_err Pp.(str "This command can only be used for unnamed theorem.") let save_anonymous ?export_seff proof save_ident = - let id,const,(cstrs,pl),do_guard,persistence,hook = proof in + let id,const,uctx,do_guard,persistence,hook = proof in check_anonymity id save_ident; - save ?export_seff save_ident const cstrs pl do_guard persistence hook + save ?export_seff save_ident const uctx do_guard persistence hook (* Admitted *) @@ -312,7 +312,7 @@ let admit (id,k,e) pl hook () = | Local, _, _ | Discharge, _, _ -> warn_let_as_axiom id in let () = assumption_message id in - Universes.register_universe_binders (ConstRef kn) (Option.default Universes.empty_binders pl); + Declare.declare_univ_binders (ConstRef kn) pl; call_hook (fun exn -> exn) hook Global (ConstRef kn) (* Starting a goal *) @@ -330,8 +330,8 @@ let get_proof proof do_guard hook opacity = let universe_proof_terminator compute_guard hook = let open Proof_global in make_terminator begin function - | Admitted (id,k,pe,(ctx,pl)) -> - admit (id,k,pe) pl (hook (Some ctx)) (); + | Admitted (id,k,pe,ctx) -> + admit (id,k,pe) (UState.universe_binders ctx) (hook (Some ctx)) (); Feedback.feedback Feedback.AddedAxiom | Proved (opaque,idopt,proof) -> let is_opaque, export_seff = match opaque with @@ -339,7 +339,7 @@ let universe_proof_terminator compute_guard hook = | Vernacexpr.Opaque -> true, false in let proof = get_proof proof compute_guard - (hook (Some (fst proof.Proof_global.universes))) is_opaque in + (hook (Some (proof.Proof_global.universes))) is_opaque in begin match idopt with | None -> save_named ~export_seff proof | Some (_,id) -> save_anonymous ~export_seff proof id @@ -417,7 +417,7 @@ let start_proof_with_initialization kind ctx decl recguard thms snl hook = | (id,(t,(_,imps)))::other_thms -> let hook ctx strength ref = let ctx = match ctx with - | None -> Evd.empty_evar_universe_context + | None -> UState.empty | Some ctx -> ctx in let other_thms_data = @@ -426,9 +426,9 @@ let start_proof_with_initialization kind ctx decl recguard thms snl hook = let body,opaq = retrieve_first_recthm ctx ref in let subst = Evd.evar_universe_context_subst ctx in let norm c = Universes.subst_opt_univs_constr subst c in - let ctx = UState.check_univ_decl ~poly:(pi2 kind) ctx decl in let body = Option.map norm body in - List.map_i (save_remaining_recthms kind norm ctx body opaq) 1 other_thms in + let uctx = UState.check_univ_decl ~poly:(pi2 kind) ctx decl in + List.map_i (save_remaining_recthms kind norm uctx body opaq) 1 other_thms in let thms_data = (strength,ref,imps)::other_thms_data in List.iter (fun (strength,ref,imps) -> maybe_declare_manual_implicits false ref imps; @@ -496,7 +496,7 @@ let save_proof ?proof = function if const_entry_type = None then user_err Pp.(str "Admitted requires an explicit statement"); let typ = Option.get const_entry_type in - let ctx = UState.const_univ_entry ~poly:(pi2 k) (fst universes) in + let ctx = UState.const_univ_entry ~poly:(pi2 k) universes in let sec_vars = if !keep_admitted_vars then const_entry_secctx else None in Admitted(id, k, (sec_vars, (typ, ctx), None), universes) | None -> @@ -518,12 +518,9 @@ let save_proof ?proof = function Some (Environ.keep_hyps env (Id.Set.union ids_typ ids_def)) | _ -> None in let decl = Proof_global.get_universe_decl () in - let evd = Evd.from_ctx universes in let poly = pi2 k in - let ctx = Evd.check_univ_decl ~poly evd decl in - let binders = if poly then Some (UState.universe_binders universes) else None in - Admitted(id,k,(sec_vars, (typ, ctx), None), - (universes, binders)) + let ctx = UState.check_univ_decl ~poly universes decl in + Admitted(id,k,(sec_vars, (typ, ctx), None), universes) in Proof_global.apply_terminator (Proof_global.get_terminator ()) pe | Vernacexpr.Proved (is_opaque,idopt) -> diff --git a/vernac/lemmas.mli b/vernac/lemmas.mli index 1b1304db5..a4854b4a6 100644 --- a/vernac/lemmas.mli +++ b/vernac/lemmas.mli @@ -56,7 +56,7 @@ val standard_proof_terminator : (** {6 ... } *) (** A hook the next three functions pass to cook_proof *) -val set_save_hook : (Proof.proof -> unit) -> unit +val set_save_hook : (Proof.t -> unit) -> unit val save_proof : ?proof:Proof_global.closed_proof -> Vernacexpr.proof_end -> unit diff --git a/vernac/locality.ml b/vernac/locality.ml index 681b1ab20..87b411625 100644 --- a/vernac/locality.ml +++ b/vernac/locality.ml @@ -6,36 +6,14 @@ (* * GNU Lesser General Public License Version 2.1 *) (************************************************************************) +open Decl_kinds + (** * Managing locality *) let local_of_bool = function - | true -> Decl_kinds.Local - | false -> Decl_kinds.Global - -(** Extracting the locality flag *) - -(* Commands which supported an inlined Local flag *) - -let warn_deprecated_local_syntax = - CWarnings.create ~name:"deprecated-local-syntax" ~category:"deprecated" - (fun () -> - Pp.strbrk "Deprecated syntax: use \"Local\" as a prefix.") - -let enforce_locality_full locality_flag local = - let local = - match locality_flag with - | Some false when local -> - CErrors.user_err Pp.(str "Cannot be simultaneously Local and Global.") - | Some true when local -> - CErrors.user_err Pp.(str "Use only prefix \"Local\".") - | None -> - if local then begin - warn_deprecated_local_syntax (); - Some true - end else - None - | Some b -> Some b in - local + | true -> Local + | false -> Global + (** Positioning locality for commands supporting discharging and export outside of modules *) @@ -48,15 +26,16 @@ let make_non_locality = function Some false -> false | _ -> true let make_locality = function Some true -> true | _ -> false -let enforce_locality locality_flag local = - make_locality (enforce_locality_full locality_flag local) +let enforce_locality_exp locality_flag discharge = + match locality_flag, discharge with + | Some b, NoDischarge -> local_of_bool b + | None, NoDischarge -> Global + | None, DoDischarge -> Discharge + | Some true, DoDischarge -> CErrors.user_err Pp.(str "Local not allowed in this case") + | Some false, DoDischarge -> CErrors.user_err Pp.(str "Global not allowed in this case") -let enforce_locality_exp locality_flag local = - match locality_flag, local with - | None, Some local -> local - | Some b, None -> local_of_bool b - | None, None -> Decl_kinds.Global - | Some _, Some _ -> CErrors.user_err Pp.(str "Local non allowed in this case") +let enforce_locality locality_flag = + make_locality locality_flag (* For commands whose default is to not discharge but to export: Global in sections forces discharge, Global not in section is the default; @@ -65,8 +44,8 @@ let enforce_locality_exp locality_flag local = let make_section_locality = function Some b -> b | None -> Lib.sections_are_opened () -let enforce_section_locality locality_flag local = - make_section_locality (enforce_locality_full locality_flag local) +let enforce_section_locality locality_flag = + make_section_locality locality_flag (** Positioning locality for commands supporting export but not discharge *) @@ -83,5 +62,5 @@ let make_module_locality = function | Some true -> true | None -> false -let enforce_module_locality locality_flag local = - make_module_locality (enforce_locality_full locality_flag local) +let enforce_module_locality locality_flag = + make_module_locality locality_flag diff --git a/vernac/locality.mli b/vernac/locality.mli index bef66d8bc..922538b23 100644 --- a/vernac/locality.mli +++ b/vernac/locality.mli @@ -8,10 +8,6 @@ (** * Managing locality *) -(** Commands which supported an inlined Local flag *) - -val enforce_locality_full : bool option -> bool -> bool option - (** * Positioning locality for commands supporting discharging and export outside of modules *) @@ -22,16 +18,15 @@ val enforce_locality_full : bool option -> bool -> bool option val make_locality : bool option -> bool val make_non_locality : bool option -> bool -val enforce_locality : bool option -> bool -> bool -val enforce_locality_exp : - bool option -> Decl_kinds.locality option -> Decl_kinds.locality +val enforce_locality_exp : bool option -> Decl_kinds.discharge -> Decl_kinds.locality +val enforce_locality : bool option -> bool (** For commands whose default is to not discharge but to export: Global in sections forces discharge, Global not in section is the default; Local in sections is the default, Local not in section forces non-export *) val make_section_locality : bool option -> bool -val enforce_section_locality : bool option -> bool -> bool +val enforce_section_locality : bool option -> bool (** * Positioning locality for commands supporting export but not discharge *) @@ -40,4 +35,4 @@ val enforce_section_locality : bool option -> bool -> bool Local in sections is the default, Local not in section forces non-export *) val make_module_locality : bool option -> bool -val enforce_module_locality : bool option -> bool -> bool +val enforce_module_locality : bool option -> bool diff --git a/vernac/obligations.ml b/vernac/obligations.ml index 1046d68f8..4f011e6ad 100644 --- a/vernac/obligations.ml +++ b/vernac/obligations.ml @@ -477,7 +477,10 @@ let declare_definition prg = let fix_exn = Hook.get get_fix_exn () in let typ = nf typ in let body = nf body in - let uvars = Univ.LSet.union (Univops.universes_of_constr typ) (Univops.universes_of_constr body) in + let env = Global.env () in + let uvars = Univ.LSet.union + (Univops.universes_of_constr env typ) + (Univops.universes_of_constr env body) in let uctx = UState.restrict prg.prg_ctx uvars in let univs = UState.check_univ_decl ~poly:(pi2 prg.prg_kind) uctx prg.prg_univdecl in let ce = definition_entry ~fix_exn ~opaque ~types:typ ~univs body in @@ -833,7 +836,7 @@ let obligation_terminator name num guard hook auto pf = let entry = Safe_typing.inline_private_constants_in_definition_entry env entry in let ty = entry.Entries.const_entry_type in let (body, cstr), () = Future.force entry.Entries.const_entry_body in - let sigma = Evd.from_ctx (fst uctx) in + let sigma = Evd.from_ctx uctx in let sigma = Evd.merge_context_set ~sideff:true Evd.univ_rigid sigma cstr in Inductiveops.control_only_guard (Global.env ()) body; (** Declare the obligation ourselves and drop the hook *) diff --git a/vernac/record.ml b/vernac/record.ml index 1d255b08e..1cdc538b5 100644 --- a/vernac/record.ml +++ b/vernac/record.ml @@ -613,7 +613,7 @@ let definition_structure (kind,cum,poly,finite,(is_coe,((loc,idstruc),pl)),ps,cf States.with_state_protection (fun () -> typecheck_params_and_fields finite (kind = Class true) idstruc poly pl s ps notations fs) () in let sign = structure_signature (fields@params) in - match kind with + let gr = match kind with | Class def -> let priorities = List.map (fun id -> {hint_priority = id; hint_pattern = None}) priorities in let gr = declare_class finite def cum pl univs (loc,idstruc) idbuild @@ -638,3 +638,6 @@ let definition_structure (kind,cum,poly,finite,(is_coe,((loc,idstruc),pl)),ps,cf idbuild implpars params arity template implfs fields is_coe (List.map (fun coe -> not (Option.is_empty coe)) coers) sign in IndRef ind + in + Declare.declare_univ_binders gr pl; + gr diff --git a/vernac/record.mli b/vernac/record.mli index 9fdd5e1c4..e632e7bbf 100644 --- a/vernac/record.mli +++ b/vernac/record.mli @@ -15,7 +15,7 @@ val primitive_flag : bool ref val definition_structure : inductive_kind * Decl_kinds.cumulative_inductive_flag * Decl_kinds.polymorphic * - Decl_kinds.recursivity_kind * ident_decl with_coercion * local_binder_expr list * + Declarations.recursivity_kind * ident_decl with_coercion * local_binder_expr list * (local_decl_expr with_instance with_priority with_notation) list * Id.t * constr_expr option -> global_reference diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml index 2f278ceb1..63768d9b8 100644 --- a/vernac/vernacentries.ml +++ b/vernac/vernacentries.ml @@ -29,6 +29,7 @@ open Redexpr open Lemmas open Misctypes open Locality +open Vernacinterp module NamedDecl = Context.Named.Declaration @@ -185,8 +186,8 @@ let print_module r = try let globdir = Nametab.locate_dir qid in match globdir with - DirModule (dirpath,(mp,_)) -> - Feedback.msg_notice (Printmod.print_module (Printmod.printable_body dirpath) mp) + DirModule { obj_dir; obj_mp; _ } -> + Feedback.msg_notice (Printmod.print_module (Printmod.printable_body obj_dir) obj_mp) | _ -> raise Not_found with Not_found -> Feedback.msg_error (str"Unknown Module " ++ pr_qualid qid) @@ -408,8 +409,8 @@ let dump_global r = (**********) (* Syntax *) -let vernac_syntax_extension locality local infix l = - let local = enforce_module_locality locality local in +let vernac_syntax_extension atts infix l = + let local = enforce_module_locality atts.locality in if infix then Metasyntax.check_infix_modifiers (snd l); Metasyntax.add_syntax_extension local l @@ -420,20 +421,20 @@ let vernac_delimiters sc = function let vernac_bind_scope sc cll = Metasyntax.add_class_scope sc (List.map scope_class_of_qualid cll) -let vernac_open_close_scope locality local (b,s) = - let local = enforce_section_locality locality local in +let vernac_open_close_scope ~atts (b,s) = + let local = enforce_section_locality atts.locality in Notation.open_close_scope (local,b,s) -let vernac_arguments_scope locality r scl = - let local = make_section_locality locality in +let vernac_arguments_scope ~atts r scl = + let local = make_section_locality atts.locality in Notation.declare_arguments_scope local (smart_global r) scl -let vernac_infix locality local = - let local = enforce_module_locality locality local in +let vernac_infix ~atts = + let local = enforce_module_locality atts.locality in Metasyntax.add_infix local (Global.env()) -let vernac_notation locality local = - let local = enforce_module_locality locality local in +let vernac_notation ~atts = + let local = enforce_module_locality atts.locality in Metasyntax.add_notation local (Global.env()) (***********) @@ -471,33 +472,33 @@ let vernac_definition_hook p = function | SubClass -> Class.add_subclass_hook p | _ -> no_hook -let vernac_definition locality p (local,k) ((loc,id as lid),pl) def = - let local = enforce_locality_exp locality local in - let hook = vernac_definition_hook p k in +let vernac_definition ~atts discharge kind ((loc,id as lid),pl) def = + let local = enforce_locality_exp atts.locality discharge in + let hook = vernac_definition_hook atts.polymorphic kind in let () = match local with | Discharge -> Dumpglob.dump_definition lid true "var" | Local | Global -> Dumpglob.dump_definition lid false "def" in (match def with | ProveBody (bl,t) -> (* local binders, typ *) - start_proof_and_print (local,p,DefinitionBody k) + start_proof_and_print (local, atts.polymorphic, DefinitionBody kind) [Some (lid,pl), (bl,t)] hook | DefineBody (bl,red_option,c,typ_opt) -> - let red_option = match red_option with + let red_option = match red_option with | None -> None | Some r -> - let sigma, env= Pfedit.get_current_context () in + let sigma, env = Pfedit.get_current_context () in Some (snd (Hook.get f_interp_redexp env sigma r)) in - do_definition id (local,p,k) pl bl red_option c typ_opt hook) + do_definition id (local, atts.polymorphic, kind) pl bl red_option c typ_opt hook) -let vernac_start_proof locality p kind l = - let local = enforce_locality_exp locality None in +let vernac_start_proof ~atts kind l = + let local = enforce_locality_exp atts.locality NoDischarge in if Dumpglob.dump () then List.iter (fun (id, _) -> match id with | Some (lid,_) -> Dumpglob.dump_definition lid false "prf" | None -> ()) l; - start_proof_and_print (local, p, Proof kind) l no_hook + start_proof_and_print (local, atts.polymorphic, Proof kind) l no_hook let vernac_end_proof ?proof = function | Admitted -> save_proof ?proof Admitted @@ -510,10 +511,10 @@ let vernac_exact_proof c = save_proof (Vernacexpr.(Proved(Opaque,None))); if not status then Feedback.feedback Feedback.AddedAxiom -let vernac_assumption locality poly (local, kind) l nl = - let local = enforce_locality_exp locality local in +let vernac_assumption ~atts discharge kind l nl = + let local = enforce_locality_exp atts.locality discharge in let global = local == Global in - let kind = local, poly, kind in + let kind = local, atts.polymorphic, kind in List.iter (fun (is_coe,(idl,c)) -> if Dumpglob.dump () then List.iter (fun (lid, _) -> @@ -553,8 +554,8 @@ let vernac_record cum k poly finite struc binders sort nameopt cfs = then the type is declared private (as per the [Private] keyword). [finite] indicates whether the type is inductive, co-inductive or neither. *) -let vernac_inductive cum poly lo finite indl = - let is_cumulative = should_treat_as_cumulative cum poly in +let vernac_inductive ~atts cum lo finite indl = + let is_cumulative = should_treat_as_cumulative cum atts.polymorphic in if Dumpglob.dump () then List.iter (fun (((coe,(lid,_)), _, _, _, cstrs), _) -> match cstrs with @@ -571,13 +572,13 @@ let vernac_inductive cum poly lo finite indl = user_err Pp.(str "The Variant keyword does not support syntax { ... }.") | [ ( id , bl , c , b, RecordDecl (oc,fs) ), [] ] -> vernac_record cum (match b with Class _ -> Class false | _ -> b) - poly finite id bl c oc fs + atts.polymorphic finite id bl c oc fs | [ ( id , bl , c , Class _, Constructors [l]), [] ] -> let f = let (coe, ((loc, id), ce)) = l in let coe' = if coe then Some true else None in (((coe', AssumExpr ((loc, Name id), ce)), None), []) - in vernac_record cum (Class true) poly finite id bl c None [f] + in vernac_record cum (Class true) atts.polymorphic finite id bl c None [f] | [ ( _ , _, _, Class _, Constructors _), [] ] -> user_err Pp.(str "Inductive classes not supported") | [ ( id , bl , c , Class _, _), _ :: _ ] -> @@ -591,19 +592,19 @@ let vernac_inductive cum poly lo finite indl = | _ -> user_err Pp.(str "Cannot handle mutually (co)inductive records.") in let indl = List.map unpack indl in - do_mutual_inductive indl is_cumulative poly lo finite + do_mutual_inductive indl is_cumulative atts.polymorphic lo finite -let vernac_fixpoint locality poly local l = - let local = enforce_locality_exp locality local in +let vernac_fixpoint ~atts discharge l = + let local = enforce_locality_exp atts.locality discharge in if Dumpglob.dump () then List.iter (fun (((lid,_), _, _, _, _), _) -> Dumpglob.dump_definition lid false "def") l; - do_fixpoint local poly l + do_fixpoint local atts.polymorphic l -let vernac_cofixpoint locality poly local l = - let local = enforce_locality_exp locality local in +let vernac_cofixpoint ~atts discharge l = + let local = enforce_locality_exp atts.locality discharge in if Dumpglob.dump () then List.iter (fun (((lid,_), _, _, _), _) -> Dumpglob.dump_definition lid false "def") l; - do_cofixpoint local poly l + do_cofixpoint local atts.polymorphic l let vernac_scheme l = if Dumpglob.dump () then @@ -621,19 +622,19 @@ let vernac_combined_scheme lid l = List.iter (fun lid -> dump_global (Misctypes.AN (Ident lid))) l); Indschemes.do_combined_scheme lid l -let vernac_universe loc poly l = - if poly && not (Lib.sections_are_opened ()) then - user_err ?loc ~hdr:"vernac_universe" +let vernac_universe ~atts l = + if atts.polymorphic && not (Lib.sections_are_opened ()) then + user_err ?loc:atts.loc ~hdr:"vernac_universe" (str"Polymorphic universes can only be declared inside sections, " ++ str "use Monomorphic Universe instead"); - do_universe poly l + do_universe atts.polymorphic l -let vernac_constraint loc poly l = - if poly && not (Lib.sections_are_opened ()) then - user_err ?loc ~hdr:"vernac_constraint" +let vernac_constraint ~atts l = + if atts.polymorphic && not (Lib.sections_are_opened ()) then + user_err ?loc:atts.loc ~hdr:"vernac_constraint" (str"Polymorphic universe constraints can only be declared" ++ str " inside sections, use Monomorphic Constraint instead"); - do_constraint poly l + do_constraint atts.polymorphic l (**********************) (* Modules *) @@ -811,32 +812,32 @@ let vernac_require from import qidl = let vernac_canonical r = Recordops.declare_canonical_structure (smart_global r) -let vernac_coercion locality poly local ref qids qidt = - let local = enforce_locality locality local in +let vernac_coercion ~atts ref qids qidt = + let local = enforce_locality atts.locality in let target = cl_of_qualid qidt in let source = cl_of_qualid qids in let ref' = smart_global ref in - Class.try_add_new_coercion_with_target ref' ~local poly ~source ~target; + Class.try_add_new_coercion_with_target ref' ~local atts.polymorphic ~source ~target; Flags.if_verbose Feedback.msg_info (pr_global ref' ++ str " is now a coercion") -let vernac_identity_coercion locality poly local id qids qidt = - let local = enforce_locality locality local in +let vernac_identity_coercion ~atts id qids qidt = + let local = enforce_locality atts.locality in let target = cl_of_qualid qidt in let source = cl_of_qualid qids in - Class.try_add_new_identity_coercion id ~local poly ~source ~target + Class.try_add_new_identity_coercion id ~local atts.polymorphic ~source ~target (* Type classes *) -let vernac_instance abst locality poly sup inst props pri = - let global = not (make_section_locality locality) in +let vernac_instance ~atts abst sup inst props pri = + let global = not (make_section_locality atts.locality) in Dumpglob.dump_constraint inst false "inst"; - ignore(Classes.new_instance ~abstract:abst ~global poly sup inst props pri) + ignore(Classes.new_instance ~abstract:abst ~global atts.polymorphic sup inst props pri) -let vernac_context poly l = - if not (Classes.context poly l) then Feedback.feedback Feedback.AddedAxiom +let vernac_context ~atts l = + if not (Classes.context atts.polymorphic l) then Feedback.feedback Feedback.AddedAxiom -let vernac_declare_instances locality insts = - let glob = not (make_section_locality locality) in +let vernac_declare_instances ~atts insts = + let glob = not (make_section_locality atts.locality) in List.iter (fun (id, info) -> Classes.existing_instance glob id (Some info)) insts let vernac_declare_class id = @@ -904,8 +905,8 @@ let vernac_remove_loadpath path = let vernac_add_ml_path isrec path = (if isrec then Mltop.add_rec_ml_dir else Mltop.add_ml_dir) (expand path) -let vernac_declare_ml_module locality l = - let local = make_locality locality in +let vernac_declare_ml_module ~atts l = + let local = make_locality atts.locality in Mltop.declare_ml_modules local (List.map expand l) let vernac_chdir = function @@ -938,25 +939,25 @@ let vernac_restore_state file = (************) (* Commands *) -let vernac_create_hintdb locality id b = - let local = make_module_locality locality in +let vernac_create_hintdb ~atts id b = + let local = make_module_locality atts.locality in Hints.create_hint_db local id full_transparent_state b -let vernac_remove_hints locality dbs ids = - let local = make_module_locality locality in +let vernac_remove_hints ~atts dbs ids = + let local = make_module_locality atts.locality in Hints.remove_hints local dbs (List.map Smartlocate.global_with_alias ids) -let vernac_hints locality poly local lb h = - let local = enforce_module_locality locality local in - Hints.add_hints local lb (Hints.interp_hints poly h) +let vernac_hints ~atts lb h = + let local = enforce_module_locality atts.locality in + Hints.add_hints local lb (Hints.interp_hints atts.polymorphic h) -let vernac_syntactic_definition locality lid x local y = +let vernac_syntactic_definition ~atts lid x y = Dumpglob.dump_definition lid false "syndef"; - let local = enforce_module_locality locality local in + let local = enforce_module_locality atts.locality in Metasyntax.add_syntactic_definition (Global.env()) (snd lid) x local y -let vernac_declare_implicits locality r l = - let local = make_section_locality locality in +let vernac_declare_implicits ~atts r l = + let local = make_section_locality atts.locality in match l with | [] -> Impargs.declare_implicits local (smart_global r) @@ -976,7 +977,7 @@ let warn_arguments_assert = (* [nargs_for_red] is the number of arguments required to trigger reduction, [args] is the main list of arguments statuses, [more_implicits] is a list of extra lists of implicit statuses *) -let vernac_arguments locality reference args more_implicits nargs_for_red flags = +let vernac_arguments ~atts reference args more_implicits nargs_for_red flags = let assert_flag = List.mem `Assert flags in let rename_flag = List.mem `Rename flags in let clear_scopes_flag = List.mem `ClearScopes flags in @@ -1184,7 +1185,7 @@ let vernac_arguments locality reference args more_implicits nargs_for_red flags (* Actions *) if renaming_specified then begin - let local = make_section_locality locality in + let local = make_section_locality atts.locality in Arguments_renaming.rename_arguments local sr names end; @@ -1194,20 +1195,20 @@ let vernac_arguments locality reference args more_implicits nargs_for_red flags with UserError _ -> Notation.find_delimiters_scope ?loc k)) scopes in - vernac_arguments_scope locality reference scopes + vernac_arguments_scope ~atts reference scopes end; if implicits_specified || clear_implicits_flag then - vernac_declare_implicits locality reference implicits; + vernac_declare_implicits ~atts reference implicits; if default_implicits_flag then - vernac_declare_implicits locality reference []; + vernac_declare_implicits ~atts reference []; if red_modifiers_specified then begin match sr with | ConstRef _ as c -> Reductionops.ReductionBehaviour.set - (make_section_locality locality) c + (make_section_locality atts.locality) c (rargs, Option.default ~-1 nargs_for_red, red_flags) | _ -> user_err (strbrk "Modifiers of the behavior of the simpl tactic "++ @@ -1235,8 +1236,8 @@ let vernac_reserve bl = Reserve.declare_reserved_type idl t) in List.iter sb_decl bl -let vernac_generalizable locality = - let local = make_non_locality locality in +let vernac_generalizable ~atts = + let local = make_non_locality atts.locality in Implicit_quantifiers.declare_generalizable local let _ = @@ -1473,8 +1474,8 @@ let _ = optread = Nativenorm.get_profiling_enabled; optwrite = Nativenorm.set_profiling_enabled } -let vernac_set_strategy locality l = - let local = make_locality locality in +let vernac_set_strategy ~atts l = + let local = make_locality atts.locality in let glob_ref r = match smart_global r with | ConstRef sp -> EvalConstRef sp @@ -1484,8 +1485,8 @@ let vernac_set_strategy locality l = let l = List.map (fun (lev,ql) -> (lev,List.map glob_ref ql)) l in Redexpr.set_strategy local l -let vernac_set_opacity locality (v,l) = - let local = make_non_locality locality in +let vernac_set_opacity ~atts (v,l) = + let local = make_non_locality atts.locality in let glob_ref r = match smart_global r with | ConstRef sp -> EvalConstRef sp @@ -1495,18 +1496,18 @@ let vernac_set_opacity locality (v,l) = let l = List.map glob_ref l in Redexpr.set_strategy local [v,l] -let vernac_set_option locality key = function - | StringValue s -> set_string_option_value_gen locality key s - | StringOptValue (Some s) -> set_string_option_value_gen locality key s - | StringOptValue None -> unset_option_value_gen locality key - | IntValue n -> set_int_option_value_gen locality key n - | BoolValue b -> set_bool_option_value_gen locality key b +let vernac_set_option ~atts key = function + | StringValue s -> set_string_option_value_gen atts.locality key s + | StringOptValue (Some s) -> set_string_option_value_gen atts.locality key s + | StringOptValue None -> unset_option_value_gen atts.locality key + | IntValue n -> set_int_option_value_gen atts.locality key n + | BoolValue b -> set_bool_option_value_gen atts.locality key b -let vernac_set_append_option locality key s = - set_string_option_append_value_gen locality key s +let vernac_set_append_option ~atts key s = + set_string_option_append_value_gen atts.locality key s -let vernac_unset_option locality key = - unset_option_value_gen locality key +let vernac_unset_option ~atts key = + unset_option_value_gen atts.locality key let vernac_add_option key lv = let f = function @@ -1547,8 +1548,8 @@ let query_command_selector ?loc = function | _ -> user_err ?loc ~hdr:"query_command_selector" (str "Query commands only support the single numbered goal selector.") -let vernac_check_may_eval ?loc redexp glopt rc = - let glopt = query_command_selector ?loc glopt in +let vernac_check_may_eval ~atts redexp glopt rc = + let glopt = query_command_selector ?loc:atts.loc glopt in let (sigma, env) = get_current_context_of_args glopt in let sigma', c = interp_open_constr env sigma rc in let c = EConstr.Unsafe.to_constr c in @@ -1582,8 +1583,8 @@ let vernac_check_may_eval ?loc redexp glopt rc = in Feedback.msg_notice (print_eval redfun env sigma' rc j) -let vernac_declare_reduction locality s r = - let local = make_locality locality in +let vernac_declare_reduction ~atts s r = + let local = make_locality atts.locality in declare_red_expr local s (snd (Hook.get f_interp_redexp (Global.env()) Evd.empty r)) (* The same but avoiding the current goal context if any *) @@ -1637,8 +1638,10 @@ let print_about_hyp_globs ?loc ref_or_by_not udecl glopt = let sigma, env = Pfedit.get_current_context () in print_about env sigma ref_or_by_not udecl - -let vernac_print ?loc env sigma = let open Feedback in function +let vernac_print ~atts env sigma = + let open Feedback in + let loc = atts.loc in + function | PrintTables -> msg_notice (print_tables ()) | PrintFullContext-> msg_notice (print_full_context_typ env sigma) | PrintSectionContext qid -> msg_notice (print_sec_context_typ env sigma qid) @@ -1747,8 +1750,8 @@ let _ = optread = (fun () -> !search_output_name_only); optwrite = (:=) search_output_name_only } -let vernac_search ?loc s gopt r = - let gopt = query_command_selector ?loc gopt in +let vernac_search ~atts s gopt r = + let gopt = query_command_selector ?loc:atts.loc gopt in let r = interp_search_restriction r in let env,gopt = match gopt with | None -> @@ -1915,7 +1918,8 @@ let vernac_load interp fname = * is the outdated/deprecated "Local" attribute of some vernacular commands * still parsed as the obsolete_locality grammar entry for retrocompatibility. * loc is the Loc.t of the vernacular command being interpreted. *) -let interp ?proof ?loc locality poly st c = +let interp ?proof ~atts ~st c = + let open Vernacinterp in vernac_pperr_endline (fun () -> str "interpreting: " ++ Ppvernac.pr_vernac c); match c with (* The below vernac are candidates for removal from the main type @@ -1954,31 +1958,33 @@ let interp ?proof ?loc locality poly st c = | VernacLocal _ -> assert false (* Syntax *) - | VernacSyntaxExtension (infix, local,sl) -> - vernac_syntax_extension locality local infix sl + | VernacSyntaxExtension (infix, sl) -> + vernac_syntax_extension atts infix sl | VernacDelimiters (sc,lr) -> vernac_delimiters sc lr | VernacBindScope (sc,rl) -> vernac_bind_scope sc rl - | VernacOpenCloseScope (local, s) -> vernac_open_close_scope locality local s - | VernacArgumentsScope (qid,scl) -> vernac_arguments_scope locality qid scl - | VernacInfix (local,mv,qid,sc) -> vernac_infix locality local mv qid sc - | VernacNotation (local,c,infpl,sc) -> - vernac_notation locality local c infpl sc + | VernacOpenCloseScope (b, s) -> vernac_open_close_scope ~atts (b,s) + | VernacArgumentsScope (qid,scl) -> vernac_arguments_scope ~atts qid scl + | VernacInfix (mv,qid,sc) -> vernac_infix ~atts mv qid sc + | VernacNotation (c,infpl,sc) -> + vernac_notation ~atts c infpl sc | VernacNotationAddFormat(n,k,v) -> Metasyntax.add_notation_extra_printing_rule n k v (* Gallina *) - | VernacDefinition (k,lid,d) -> vernac_definition locality poly k lid d - | VernacStartTheoremProof (k,l) -> vernac_start_proof locality poly k l + | VernacDefinition ((discharge,kind),lid,d) -> + vernac_definition ~atts discharge kind lid d + | VernacStartTheoremProof (k,l) -> vernac_start_proof ~atts k l | VernacEndProof e -> vernac_end_proof ?proof e | VernacExactProof c -> vernac_exact_proof c - | VernacAssumption (stre,nl,l) -> vernac_assumption locality poly stre l nl - | VernacInductive (cum, priv,finite,l) -> vernac_inductive cum poly priv finite l - | VernacFixpoint (local, l) -> vernac_fixpoint locality poly local l - | VernacCoFixpoint (local, l) -> vernac_cofixpoint locality poly local l + | VernacAssumption ((discharge,kind),nl,l) -> + vernac_assumption ~atts discharge kind l nl + | VernacInductive (cum, priv,finite,l) -> vernac_inductive ~atts cum priv finite l + | VernacFixpoint (discharge, l) -> vernac_fixpoint ~atts discharge l + | VernacCoFixpoint (discharge, l) -> vernac_cofixpoint ~atts discharge l | VernacScheme l -> vernac_scheme l | VernacCombinedScheme (id, l) -> vernac_combined_scheme id l - | VernacUniverse l -> vernac_universe loc poly l - | VernacConstraint l -> vernac_constraint loc poly l + | VernacUniverse l -> vernac_universe ~atts l + | VernacConstraint l -> vernac_constraint ~atts l (* Modules *) | VernacDeclareModule (export,lid,bl,mtyo) -> @@ -1999,15 +2005,15 @@ let interp ?proof ?loc locality poly st c = | VernacRequire (from, export, qidl) -> vernac_require from export qidl | VernacImport (export,qidl) -> vernac_import export qidl | VernacCanonical qid -> vernac_canonical qid - | VernacCoercion (local,r,s,t) -> vernac_coercion locality poly local r s t - | VernacIdentityCoercion (local,(_,id),s,t) -> - vernac_identity_coercion locality poly local id s t + | VernacCoercion (r,s,t) -> vernac_coercion ~atts r s t + | VernacIdentityCoercion ((_,id),s,t) -> + vernac_identity_coercion ~atts id s t (* Type classes *) | VernacInstance (abst, sup, inst, props, info) -> - vernac_instance abst locality poly sup inst props info - | VernacContext sup -> vernac_context poly sup - | VernacDeclareInstances insts -> vernac_declare_instances locality insts + vernac_instance ~atts abst sup inst props info + | VernacContext sup -> vernac_context ~atts sup + | VernacDeclareInstances insts -> vernac_declare_instances ~atts insts | VernacDeclareClass id -> vernac_declare_class id (* Solving *) @@ -2017,7 +2023,7 @@ let interp ?proof ?loc locality poly st c = | VernacAddLoadPath (isrec,s,alias) -> vernac_add_loadpath isrec s alias | VernacRemoveLoadPath s -> vernac_remove_loadpath s | VernacAddMLPath (isrec,s) -> vernac_add_ml_path isrec s - | VernacDeclareMLModule l -> vernac_declare_ml_module locality l + | VernacDeclareMLModule l -> vernac_declare_ml_module ~atts l | VernacChdir s -> vernac_chdir s (* State management *) @@ -2025,40 +2031,40 @@ let interp ?proof ?loc locality poly st c = | VernacRestoreState s -> vernac_restore_state s (* Commands *) - | VernacCreateHintDb (dbname,b) -> vernac_create_hintdb locality dbname b - | VernacRemoveHints (dbnames,ids) -> vernac_remove_hints locality dbnames ids - | VernacHints (local,dbnames,hints) -> - vernac_hints locality poly local dbnames hints - | VernacSyntacticDefinition (id,c,local,b) -> - vernac_syntactic_definition locality id c local b + | VernacCreateHintDb (dbname,b) -> vernac_create_hintdb ~atts dbname b + | VernacRemoveHints (dbnames,ids) -> vernac_remove_hints ~atts dbnames ids + | VernacHints (dbnames,hints) -> + vernac_hints ~atts dbnames hints + | VernacSyntacticDefinition (id,c,b) -> + vernac_syntactic_definition ~atts id c b | VernacDeclareImplicits (qid,l) -> - vernac_declare_implicits locality qid l + vernac_declare_implicits ~atts qid l | VernacArguments (qid, args, more_implicits, nargs, flags) -> - vernac_arguments locality qid args more_implicits nargs flags + vernac_arguments ~atts qid args more_implicits nargs flags | VernacReserve bl -> vernac_reserve bl - | VernacGeneralizable gen -> vernac_generalizable locality gen - | VernacSetOpacity qidl -> vernac_set_opacity locality qidl - | VernacSetStrategy l -> vernac_set_strategy locality l - | VernacSetOption (key,v) -> vernac_set_option locality key v - | VernacSetAppendOption (key,v) -> vernac_set_append_option locality key v - | VernacUnsetOption key -> vernac_unset_option locality key + | VernacGeneralizable gen -> vernac_generalizable ~atts gen + | VernacSetOpacity qidl -> vernac_set_opacity ~atts qidl + | VernacSetStrategy l -> vernac_set_strategy ~atts l + | VernacSetOption (key,v) -> vernac_set_option ~atts key v + | VernacSetAppendOption (key,v) -> vernac_set_append_option ~atts key v + | VernacUnsetOption key -> vernac_unset_option ~atts key | VernacRemoveOption (key,v) -> vernac_remove_option key v | VernacAddOption (key,v) -> vernac_add_option key v | VernacMemOption (key,v) -> vernac_mem_option key v | VernacPrintOption key -> vernac_print_option key - | VernacCheckMayEval (r,g,c) -> vernac_check_may_eval ?loc r g c - | VernacDeclareReduction (s,r) -> vernac_declare_reduction locality s r + | VernacCheckMayEval (r,g,c) -> vernac_check_may_eval ~atts r g c + | VernacDeclareReduction (s,r) -> vernac_declare_reduction ~atts s r | VernacGlobalCheck c -> vernac_global_check c | VernacPrint p -> let sigma, env = Pfedit.get_current_context () in - vernac_print ?loc env sigma p - | VernacSearch (s,g,r) -> vernac_search ?loc s g r + vernac_print ~atts env sigma p + | VernacSearch (s,g,r) -> vernac_search ~atts s g r | VernacLocate l -> vernac_locate l | VernacRegister (id, r) -> vernac_register id r | VernacComments l -> Flags.if_verbose Feedback.msg_info (str "Comments ok\n") (* Proof management *) - | VernacGoal t -> vernac_start_proof locality poly Theorem [None,([],t)] + | VernacGoal t -> vernac_start_proof ~atts Theorem [None,([],t)] | VernacFocus n -> vernac_focus n | VernacUnfocus -> vernac_unfocus () | VernacUnfocused -> vernac_unfocused () @@ -2071,7 +2077,7 @@ let interp ?proof ?loc locality poly st c = let using = Option.append using (Proof_using.get_default_proof_using ()) in let tacs = if Option.is_empty tac then "tac:no" else "tac:yes" in let usings = if Option.is_empty using then "using:no" else "using:yes" in - Aux_file.record_in_aux_at ?loc "VernacProof" (tacs^" "^usings); + Aux_file.record_in_aux_at ?loc:atts.loc "VernacProof" (tacs^" "^usings); Option.iter vernac_set_end_tac tac; Option.iter vernac_set_used_variables using | VernacProofMode mn -> Proof_global.set_proof_mode mn [@ocaml.warning "-3"] @@ -2079,7 +2085,7 @@ let interp ?proof ?loc locality poly st c = (* Extensions *) | VernacExtend (opn,args) -> (* XXX: Here we are returning the state! :) *) - let _st : Vernacstate.t = Vernacinterp.call ?locality ?loc (opn,args) ~st in + let _st : Vernacstate.t = Vernacinterp.call ~atts opn args ~st in () (* Vernaculars that take a locality flag *) @@ -2111,7 +2117,7 @@ let check_vernac_supports_polymorphism c p = | None, _ -> () | Some _, ( VernacDefinition _ | VernacFixpoint _ | VernacCoFixpoint _ - | VernacAssumption _ | VernacInductive _ + | VernacAssumption _ | VernacInductive _ | VernacStartTheoremProof _ | VernacCoercion _ | VernacIdentityCoercion _ | VernacInstance _ | VernacDeclareInstances _ @@ -2120,7 +2126,7 @@ let check_vernac_supports_polymorphism c p = | Some _, _ -> user_err Pp.(str "This command does not support Polymorphism") let enforce_polymorphism = function - | None -> Flags.is_universe_polymorphism () + | None -> Flags.is_universe_polymorphism () | Some b -> Flags.make_polymorphic_flag b; b (** A global default timeout, controlled by option "Set Default Timeout n". @@ -2188,42 +2194,57 @@ let with_fail st b f = | _ -> assert false end -let interp ?(verbosely=true) ?proof st (loc,c) = +let interp ?(verbosely=true) ?proof ~st (loc,c) = let orig_program_mode = Flags.is_program_mode () in - let rec aux ?locality ?polymorphism isprogcmd = function - - | VernacProgram c when not isprogcmd -> aux ?locality ?polymorphism true c - | VernacProgram _ -> user_err Pp.(str "Program mode specified twice") - | VernacLocal (b, c) when Option.is_empty locality -> - aux ~locality:b ?polymorphism isprogcmd c - | VernacPolymorphic (b, c) when polymorphism = None -> - aux ?locality ~polymorphism:b isprogcmd c - | VernacPolymorphic (b, c) -> user_err Pp.(str "Polymorphism specified twice") - | VernacLocal _ -> user_err Pp.(str "Locality specified twice") + let rec aux ?polymorphism ~atts isprogcmd = function + + | VernacProgram c when not isprogcmd -> + aux ?polymorphism ~atts true c + + | VernacProgram _ -> + user_err Pp.(str "Program mode specified twice") + + | VernacPolymorphic (b, c) when polymorphism = None -> + aux ~polymorphism:b ~atts:atts isprogcmd c + + | VernacPolymorphic (b, c) -> + user_err Pp.(str "Polymorphism specified twice") + + | VernacLocal (b, c) when Option.is_empty atts.locality -> + aux ?polymorphism ~atts:{atts with locality = Some b} isprogcmd c + + | VernacLocal _ -> + user_err Pp.(str "Locality specified twice") + | VernacFail v -> - with_fail st true (fun () -> aux ?locality ?polymorphism isprogcmd v) + with_fail st true (fun () -> aux ?polymorphism ~atts isprogcmd v) + | VernacTimeout (n,v) -> - current_timeout := Some n; - aux ?locality ?polymorphism isprogcmd v + current_timeout := Some n; + aux ?polymorphism ~atts isprogcmd v + | VernacRedirect (s, (_,v)) -> - Topfmt.with_output_to_file s (aux ?locality ?polymorphism isprogcmd) v + Topfmt.with_output_to_file s (aux ?polymorphism ~atts isprogcmd) v + | VernacTime (_,v) -> - System.with_time !Flags.time - (aux ?locality ?polymorphism isprogcmd) v; - | VernacLoad (_,fname) -> vernac_load (aux false) fname - | c -> - check_vernac_supports_locality c locality; - check_vernac_supports_polymorphism c polymorphism; - let poly = enforce_polymorphism polymorphism in - Obligations.set_program_mode isprogcmd; - try - vernac_timeout begin fun () -> + System.with_time !Flags.time (aux ?polymorphism ~atts isprogcmd) v; + + | VernacLoad (_,fname) -> vernac_load (aux ?polymorphism ~atts false) fname + + | c -> + check_vernac_supports_locality c atts.locality; + check_vernac_supports_polymorphism c polymorphism; + let polymorphic = enforce_polymorphism polymorphism in + Obligations.set_program_mode isprogcmd; + try + vernac_timeout begin fun () -> + let atts = { atts with polymorphic } in if verbosely - then Flags.verbosely (interp ?proof ?loc locality poly st) c - else Flags.silently (interp ?proof ?loc locality poly st) c; + then Flags.verbosely (interp ?proof ~atts ~st) c + else Flags.silently (interp ?proof ~atts ~st) c; if orig_program_mode || not !Flags.program_mode || isprogcmd then Flags.program_mode := orig_program_mode; - ignore (Flags.use_polymorphic_flag ()) + ignore (Flags.use_polymorphic_flag ()) end with | reraise when @@ -2238,12 +2259,19 @@ let interp ?(verbosely=true) ?proof st (loc,c) = ignore (Flags.use_polymorphic_flag ()); iraise e in - if verbosely then Flags.verbosely (aux false) c - else aux false c + let atts = { loc; locality = None; polymorphic = false; } in + if verbosely + then Flags.verbosely (aux ~atts orig_program_mode) c + else aux ~atts orig_program_mode c (* XXX: There is a bug here in case of an exception, see @gares comments on the PR *) -let interp ?verbosely ?proof st cmd = +let interp ?verbosely ?proof ~st cmd = Vernacstate.unfreeze_interp_state st; - interp ?verbosely ?proof st cmd; - Vernacstate.freeze_interp_state `No + try + interp ?verbosely ?proof ~st cmd; + Vernacstate.freeze_interp_state `No + with exn -> + let exn = CErrors.push exn in + Vernacstate.invalidate_cache (); + iraise exn diff --git a/vernac/vernacentries.mli b/vernac/vernacentries.mli index 67001bc24..a559912a0 100644 --- a/vernac/vernacentries.mli +++ b/vernac/vernacentries.mli @@ -18,7 +18,7 @@ val vernac_require : val interp : ?verbosely:bool -> ?proof:Proof_global.closed_proof -> - Vernacstate.t -> Vernacexpr.vernac_expr Loc.located -> Vernacstate.t + st:Vernacstate.t -> Vernacexpr.vernac_expr Loc.located -> Vernacstate.t (** Prepare a "match" template for a given inductive type. For each branch of the match, we list the constructor name diff --git a/vernac/vernacinterp.ml b/vernac/vernacinterp.ml index 47dec1958..c0b93c163 100644 --- a/vernac/vernacinterp.ml +++ b/vernac/vernacinterp.ml @@ -15,17 +15,17 @@ type deprecation = bool type atts = { loc : Loc.t option; locality : bool option; + polymorphic : bool; } -type vernac_command = - Genarg.raw_generic_argument list -> - atts:atts -> st:Vernacstate.t -> - Vernacstate.t +type 'a vernac_command = 'a -> atts:atts -> st:Vernacstate.t -> Vernacstate.t + +type plugin_args = Genarg.raw_generic_argument list (* Table of vernac entries *) let vernac_tab = (Hashtbl.create 211 : - (Vernacexpr.extend_name, deprecation * vernac_command) Hashtbl.t) + (Vernacexpr.extend_name, deprecation * plugin_args vernac_command) Hashtbl.t) let vinterp_add depr s f = try @@ -58,7 +58,7 @@ let warn_deprecated_command = (* Interpretation of a vernac command *) -let call ?locality ?loc (opn,converted_args) = +let call opn converted_args ~atts ~st = let phase = ref "Looking up command" in try let depr, callback = vinterp_map opn in @@ -74,8 +74,7 @@ let call ?locality ?loc (opn,converted_args) = phase := "Checking arguments"; let hunk = callback converted_args in phase := "Executing command"; - let atts = { loc; locality } in - hunk ~atts + hunk ~atts ~st with | Drop -> raise Drop | reraise -> diff --git a/vernac/vernacinterp.mli b/vernac/vernacinterp.mli index 602ccba15..ab3d4bfc5 100644 --- a/vernac/vernacinterp.mli +++ b/vernac/vernacinterp.mli @@ -13,19 +13,15 @@ type deprecation = bool type atts = { loc : Loc.t option; locality : bool option; + polymorphic : bool; } -type vernac_command = - Genarg.raw_generic_argument list -> - atts:atts -> st:Vernacstate.t -> - Vernacstate.t +type 'a vernac_command = 'a -> atts:atts -> st:Vernacstate.t -> Vernacstate.t -val vinterp_add : deprecation -> Vernacexpr.extend_name -> vernac_command -> unit - -val overwriting_vinterp_add : Vernacexpr.extend_name -> vernac_command -> unit +type plugin_args = Genarg.raw_generic_argument list val vinterp_init : unit -> unit +val vinterp_add : deprecation -> Vernacexpr.extend_name -> plugin_args vernac_command -> unit +val overwriting_vinterp_add : Vernacexpr.extend_name -> plugin_args vernac_command -> unit -val call : ?locality:bool -> ?loc:Loc.t -> - Vernacexpr.extend_name * Genarg.raw_generic_argument list -> - st:Vernacstate.t -> Vernacstate.t +val call : Vernacexpr.extend_name -> plugin_args -> atts:atts -> st:Vernacstate.t -> Vernacstate.t diff --git a/vernac/vernacstate.ml b/vernac/vernacstate.ml index eb1359d52..4980333b5 100644 --- a/vernac/vernacstate.ml +++ b/vernac/vernacstate.ml @@ -8,22 +8,34 @@ type t = { system : States.state; (* summary + libstack *) - proof : Proof_global.state; (* proof state *) + proof : Proof_global.t; (* proof state *) shallow : bool (* is the state trimmed down (libstack) *) } -let s_cache = ref (States.freeze ~marshallable:`No) -let s_proof = ref (Proof_global.freeze ~marshallable:`No) +let s_cache = ref None +let s_proof = ref None let invalidate_cache () = - s_cache := Obj.magic 0; - s_proof := Obj.magic 0 + s_cache := None; + s_proof := None + +let update_cache rf v = + rf := Some v; v + +let do_if_not_cached rf f v = + match !rf with + | None -> + rf := Some v; f v + | Some vc when vc != v -> + rf := Some v; f v + | Some _ -> + () let freeze_interp_state marshallable = - { system = (s_cache := States.freeze ~marshallable; !s_cache); - proof = (s_proof := Proof_global.freeze ~marshallable; !s_proof); + { system = update_cache s_cache (States.freeze ~marshallable); + proof = update_cache s_proof (Proof_global.freeze ~marshallable); shallow = marshallable = `Shallow } let unfreeze_interp_state { system; proof } = - if (!s_cache != system) then (s_cache := system; States.unfreeze system); - if (!s_proof != proof) then (s_proof := proof; Proof_global.unfreeze proof) + do_if_not_cached s_cache States.unfreeze system; + do_if_not_cached s_proof Proof_global.unfreeze proof diff --git a/vernac/vernacstate.mli b/vernac/vernacstate.mli index bcfa49aa3..3ed27ddb7 100644 --- a/vernac/vernacstate.mli +++ b/vernac/vernacstate.mli @@ -8,7 +8,7 @@ type t = { system : States.state; (* summary + libstack *) - proof : Proof_global.state; (* proof state *) + proof : Proof_global.t; (* proof state *) shallow : bool (* is the state trimmed down (libstack) *) } |