From b2f2727670853183bfbcbafb9dc19f0f71494a7b Mon Sep 17 00:00:00 2001 From: gareuselesinge Date: Thu, 8 Aug 2013 18:51:35 +0000 Subject: State Transaction Machine The process_transaction function adds a new edge to the Dag without executing the transaction (when possible). The observe id function runs the transactions necessary to reach to the state id. Transaction being on a merged branch are not executed but stored into a future. The finish function calls observe on the tip of the current branch. Imperative modifications to the environment made by some tactics are now explicitly declared by the tactic and modeled as let-in/beta-redexes at the root of the proof term. An example is the abstract tactic. This is the work described in the Coq Workshop 2012 paper. Coq is compile with thread support from now on. git-svn-id: svn+ssh://scm.gforge.inria.fr/svn/coq/trunk@16674 85f007b7-540e-0410-9357-904b9bb8a0f7 --- Makefile.build | 4 +- checker/check.mllib | 1 + checker/cic.mli | 4 +- checker/declarations.ml | 7 +- dev/printers.mllib | 2 + intf/vernacexpr.mli | 39 ++ kernel/cooking.ml | 37 +- kernel/cooking.mli | 7 +- kernel/declarations.mli | 8 +- kernel/declareops.ml | 39 +- kernel/declareops.mli | 14 + kernel/entries.mli | 3 +- kernel/mod_typing.ml | 18 +- kernel/modops.ml | 29 + kernel/modops.mli | 4 + kernel/safe_typing.ml | 91 ++- kernel/safe_typing.mli | 11 + kernel/term_typing.ml | 199 ++++-- kernel/term_typing.mli | 19 +- lib/flags.ml | 2 + lib/flags.mli | 2 + lib/vcs.ml | 13 +- lib/vcs.mli | 4 +- library/global.ml | 12 +- library/global.mli | 2 + library/lib.ml | 80 +-- library/lib.mli | 25 - library/library.ml | 11 +- library/nametab.ml | 4 +- library/states.ml | 11 - library/states.mli | 12 - parsing/g_vernac.ml4 | 8 + plugins/decl_mode/decl_mode.ml | 8 +- plugins/decl_mode/decl_mode.mli | 4 +- plugins/decl_mode/decl_proof_instr.ml | 43 +- plugins/decl_mode/g_decl_mode.ml4 | 14 +- plugins/extraction/extraction.ml | 4 +- plugins/funind/functional_principles_types.ml | 8 +- plugins/funind/invfun.ml | 17 +- plugins/funind/recdef.ml | 15 +- plugins/setoid_ring/newring.ml4 | 18 +- pretyping/evd.ml | 7 +- pretyping/evd.mli | 6 +- printing/ppvernac.ml | 15 +- printing/prettyp.ml | 7 +- printing/printer.ml | 10 +- proofs/clenvtac.ml | 2 +- proofs/goal.ml | 88 +-- proofs/goal.mli | 14 +- proofs/logic.ml | 26 +- proofs/pfedit.ml | 91 +-- proofs/pfedit.mli | 38 +- proofs/proof.ml | 278 ++------ proofs/proof.mli | 58 +- proofs/proof_global.ml | 257 ++++--- proofs/proof_global.mli | 49 +- proofs/proofview.ml | 75 +- proofs/proofview.mli | 9 +- proofs/refiner.ml | 57 +- proofs/refiner.mli | 8 +- proofs/tacmach.ml | 3 +- proofs/tacmach.mli | 10 +- tactics/auto.ml | 2 +- tactics/class_tactics.ml4 | 41 +- tactics/elimschemes.ml | 24 +- tactics/eqschemes.ml | 29 +- tactics/eqschemes.mli | 6 +- tactics/equality.ml | 49 +- tactics/extratactics.ml4 | 10 +- tactics/leminv.ml | 9 +- tactics/rewrite.ml4 | 45 +- tactics/tacinterp.ml | 8 +- tactics/tactics.ml | 9 + tactics/tactics.mli | 3 + tools/coqmktop.ml | 2 +- tools/fake_ide.ml | 36 +- toplevel/auto_ind_decl.ml | 111 +-- toplevel/auto_ind_decl.mli | 8 +- toplevel/autoinstance.ml | 6 +- toplevel/backtrack.ml | 234 ------- toplevel/backtrack.mli | 101 --- toplevel/class.ml | 3 +- toplevel/classes.ml | 5 +- toplevel/classes.mli | 2 +- toplevel/command.ml | 28 +- toplevel/command.mli | 2 +- toplevel/coqtop.ml | 7 +- toplevel/ide_slave.ml | 31 +- toplevel/ind_tables.ml | 30 +- toplevel/ind_tables.mli | 15 +- toplevel/indschemes.ml | 9 +- toplevel/lemmas.ml | 69 +- toplevel/lemmas.mli | 14 +- toplevel/obligations.ml | 63 +- toplevel/obligations.mli | 2 +- toplevel/record.ml | 9 +- toplevel/stm.ml | 952 ++++++++++++++++++++++++++ toplevel/stm.mli | 27 + toplevel/toplevel.ml | 26 +- toplevel/toplevel.mllib | 6 +- toplevel/vernac.ml | 130 +--- toplevel/vernac.mli | 8 +- toplevel/vernac_classifier.ml | 164 +++++ toplevel/vernac_classifier.mli | 20 + toplevel/vernacentries.ml | 385 +++++------ toplevel/vernacentries.mli | 11 +- toplevel/whelp.ml4 | 4 +- 107 files changed, 2801 insertions(+), 1915 deletions(-) delete mode 100644 toplevel/backtrack.ml delete mode 100644 toplevel/backtrack.mli create mode 100644 toplevel/stm.ml create mode 100644 toplevel/stm.mli create mode 100644 toplevel/vernac_classifier.ml create mode 100644 toplevel/vernac_classifier.mli diff --git a/Makefile.build b/Makefile.build index b68a02cf6..d0007f618 100644 --- a/Makefile.build +++ b/Makefile.build @@ -106,8 +106,8 @@ MLINCLUDES=$(LOCALINCLUDES) -I $(MYCAMLP4LIB) OCAMLC := $(TYPEREX) $(OCAMLC) $(CAMLFLAGS) OCAMLOPT := $(TYPEREX) $(OCAMLOPT) $(CAMLFLAGS) -BYTEFLAGS=$(CAMLDEBUG) $(USERFLAGS) -OPTFLAGS=$(CAMLDEBUGOPT) $(CAMLTIMEPROF) $(USERFLAGS) +BYTEFLAGS=$(CAMLDEBUG) $(USERFLAGS) -thread +OPTFLAGS=$(CAMLDEBUGOPT) $(CAMLTIMEPROF) $(USERFLAGS) -thread DEPFLAGS= -slash $(LOCALINCLUDES) define bestocaml diff --git a/checker/check.mllib b/checker/check.mllib index cd858000a..24479fb8a 100644 --- a/checker/check.mllib +++ b/checker/check.mllib @@ -21,6 +21,7 @@ CList CString CArray Util +Future CUnix System Envars diff --git a/checker/cic.mli b/checker/cic.mli index d55a9d5c0..fd3351674 100644 --- a/checker/cic.mli +++ b/checker/cic.mli @@ -181,14 +181,14 @@ type inline = int option type constant_def = | Undef of inline | Def of constr_substituted - | OpaqueDef of lazy_constr + | OpaqueDef of lazy_constr Future.computation type constant_body = { const_hyps : section_context; (** New: younger hyp at top *) const_body : constant_def; const_type : constant_type; const_body_code : to_patch_substituted; - const_constraints : Univ.constraints; + const_constraints : Univ.constraints Future.computation; const_native_name : native_name ref; const_inline_code : bool } diff --git a/checker/declarations.ml b/checker/declarations.ml index 58a5f2602..0981cfd1a 100644 --- a/checker/declarations.ml +++ b/checker/declarations.ml @@ -434,12 +434,15 @@ let force_lazy_constr = function let subst_constant_def sub = function | Undef inl -> Undef inl | Def c -> Def (subst_constr_subst sub c) - | OpaqueDef lc -> OpaqueDef (subst_lazy_constr sub lc) + | OpaqueDef lc -> + OpaqueDef (Future.from_val (subst_lazy_constr sub (Future.join lc))) + +type constant_constraints = Univ.constraints Future.computation let body_of_constant cb = match cb.const_body with | Undef _ -> None | Def c -> Some (force_constr c) - | OpaqueDef c -> Some (force_lazy_constr c) + | OpaqueDef c -> Some (force_lazy_constr (Future.join c)) let constant_has_body cb = match cb.const_body with | Undef _ -> false diff --git a/dev/printers.mllib b/dev/printers.mllib index 37ee00517..5d79e4ee9 100644 --- a/dev/printers.mllib +++ b/dev/printers.mllib @@ -123,6 +123,8 @@ Unification Cases Pretyping Declaremods +Library +States Genprint Tok diff --git a/intf/vernacexpr.mli b/intf/vernacexpr.mli index d19857ee1..2f8ab3813 100644 --- a/intf/vernacexpr.mli +++ b/intf/vernacexpr.mli @@ -8,6 +8,8 @@ open Loc open Pp +open Util +open Stateid open Names open Tacexpr open Misctypes @@ -214,6 +216,13 @@ type bullet = | Star | Plus +(** {6 Types concerning Stm} *) +type 'a stm_vernac = + | JoinDocument + | Finish + | Observe of Stateid.state_id + | Command of 'a (* An out of flow command not to be recorded by Stm *) + (** {6 Types concerning the module layer} *) (** Rigid / flexible module signature *) @@ -372,6 +381,9 @@ type vernac_expr = | VernacComments of comment list | VernacNop + (* Stm backdoor *) + | VernacStm of vernac_expr stm_vernac + (* Proof management *) | VernacGoal of constr_expr | VernacAbort of lident option @@ -401,3 +413,30 @@ type vernac_expr = | VernacLocal of bool * vernac_expr and located_vernac_expr = Loc.t * vernac_expr + +(* A vernac classifier has to tell if a command: + vernac_when: has to be executed now (alters the parser) or later + vernac_type: if it is starts, ends, continues a proof or + alters the global state or is a control command like BackTo or is + a query like Check *) +type vernac_type = + | VtStartProof of vernac_start + | VtSideff + | VtQed of vernac_qed_type + | VtProofStep + | VtQuery of vernac_part_of_script + | VtStm of vernac_control * vernac_part_of_script + | VtUnknown +and vernac_qed_type = KeepProof | DropProof (* Qed, Admitted/Abort *) +and vernac_start = string * Id.t list +and vernac_is_alias = bool +and vernac_part_of_script = bool +and vernac_control = + | VtFinish + | VtJoinDocument + | VtObserve of state_id + | VtBack of state_id +type vernac_when = + | VtNow + | VtLater +type vernac_classification = vernac_type * vernac_when diff --git a/kernel/cooking.ml b/kernel/cooking.ml index bb29b9645..4c2f22199 100644 --- a/kernel/cooking.ml +++ b/kernel/cooking.ml @@ -42,11 +42,7 @@ type my_global_reference = | IndRef of inductive | ConstructRef of constructor -let cache = (Hashtbl.create 13 : (my_global_reference, constr) Hashtbl.t) - -let clear_cooking_sharing () = Hashtbl.clear cache - -let share r (cstl,knl) = +let share cache r (cstl,knl) = try Hashtbl.find cache r with Not_found -> let f,l = @@ -59,13 +55,12 @@ let share r (cstl,knl) = mkConst (pop_con cst), Cmap.find cst cstl in let c = mkApp (f, Array.map mkVar l) in Hashtbl.add cache r c; - (* has raised Not_found if not in work_list *) c -let update_case_info ci modlist = +let update_case_info cache ci modlist = try let ind, n = - match kind_of_term (share (IndRef ci.ci_ind) modlist) with + match kind_of_term (share cache (IndRef ci.ci_ind) modlist) with | App (f,l) -> (destInd f, Array.length l) | Ind ind -> ind, 0 | _ -> assert false in @@ -78,7 +73,9 @@ let empty_modlist = (Cmap.empty, Mindmap.empty) let is_empty_modlist (cm, mm) = Cmap.is_empty cm && Mindmap.is_empty mm -let expmod_constr modlist c = +let expmod_constr cache modlist c = + let share = share cache in + let update_case_info = update_case_info cache in let rec substrec c = match kind_of_term c with | Case (ci,p,t,br) -> @@ -122,25 +119,27 @@ type recipe = { type inline = bool type result = - constant_def * constant_type * Univ.constraints * inline + constant_def * constant_type * constant_constraints * inline * Context.section_context option let on_body f = function - | Undef inl -> Undef inl + | Undef _ as x -> x | Def cs -> Def (Lazyconstr.from_val (f (Lazyconstr.force cs))) | OpaqueDef lc -> - OpaqueDef (Lazyconstr.opaque_from_val (f (Lazyconstr.force_opaque lc))) + OpaqueDef (Future.chain ~id:"Cooking.on_body" ~pure:true lc (fun lc -> + (Lazyconstr.opaque_from_val (f (Lazyconstr.force_opaque lc))))) let constr_of_def = function | Undef _ -> assert false | Def cs -> Lazyconstr.force cs - | OpaqueDef lc -> Lazyconstr.force_opaque lc + | OpaqueDef lc -> Lazyconstr.force_opaque (Future.force lc) let cook_constant env r = + let cache = Hashtbl.create 13 in let cb = r.d_from in - let hyps = Context.map_named_context (expmod_constr r.d_modlist) r.d_abstract in + let hyps = Context.map_named_context (expmod_constr cache r.d_modlist) r.d_abstract in let body = on_body - (fun c -> abstract_constant_body (expmod_constr r.d_modlist c) hyps) + (fun c -> abstract_constant_body (expmod_constr cache r.d_modlist c) hyps) cb.const_body in let const_hyps = @@ -149,12 +148,16 @@ let cook_constant env r = hyps ~init:cb.const_hyps in let typ = match cb.const_type with | NonPolymorphicType t -> - let typ = abstract_constant_type (expmod_constr r.d_modlist t) hyps in + let typ = + abstract_constant_type (expmod_constr cache r.d_modlist t) hyps in NonPolymorphicType typ | PolymorphicArity (ctx,s) -> let t = mkArity (ctx,Type s.poly_level) in - let typ = abstract_constant_type (expmod_constr r.d_modlist t) hyps in + let typ = + abstract_constant_type (expmod_constr cache r.d_modlist t) hyps in let j = make_judge (constr_of_def body) typ in Typeops.make_polymorphic_if_constant_for_ind env j in (body, typ, cb.const_constraints, cb.const_inline_code, Some const_hyps) + +let expmod_constr modlist c = expmod_constr (Hashtbl.create 13) modlist c diff --git a/kernel/cooking.mli b/kernel/cooking.mli index 8d046a12e..2d6e53407 100644 --- a/kernel/cooking.mli +++ b/kernel/cooking.mli @@ -24,17 +24,12 @@ type recipe = { type inline = bool type result = - constant_def * constant_type * constraints * inline + constant_def * constant_type * constant_constraints * inline * Context.section_context option val cook_constant : env -> recipe -> result - (** {6 Utility functions used in module [Discharge]. } *) val expmod_constr : work_list -> constr -> constr -val clear_cooking_sharing : unit -> unit - - - diff --git a/kernel/declarations.mli b/kernel/declarations.mli index 43b908e1f..5cb406ffa 100644 --- a/kernel/declarations.mli +++ b/kernel/declarations.mli @@ -38,7 +38,7 @@ type inline = int option type constant_def = | Undef of inline | Def of Lazyconstr.constr_substituted - | OpaqueDef of Lazyconstr.lazy_constr + | OpaqueDef of Lazyconstr.lazy_constr Future.computation (** Linking information for the native compiler. The boolean flag indicates if the term is protected by a lazy tag *) @@ -48,15 +48,19 @@ type native_name = | LinkedInteractive of string * bool | NotLinked +type constant_constraints = Univ.constraints Future.computation + type constant_body = { const_hyps : Context.section_context; (** New: younger hyp at top *) const_body : constant_def; const_type : constant_type; const_body_code : Cemitcodes.to_patch_substituted; - const_constraints : Univ.constraints; + const_constraints : constant_constraints; const_native_name : native_name ref; const_inline_code : bool } +type side_effect = NewConstant of constant * constant_body + (** {6 Representation of mutual inductive types in the kernel } *) type recarg = diff --git a/kernel/declareops.ml b/kernel/declareops.ml index 64496033a..e597ea9a9 100644 --- a/kernel/declareops.ml +++ b/kernel/declareops.ml @@ -19,7 +19,7 @@ open Util let body_of_constant cb = match cb.const_body with | Undef _ -> None | Def c -> Some (Lazyconstr.force c) - | OpaqueDef lc -> Some (Lazyconstr.force_opaque lc) + | OpaqueDef lc -> Some (Lazyconstr.force_opaque (Future.force lc)) let constant_has_body cb = match cb.const_body with | Undef _ -> false @@ -50,7 +50,8 @@ let subst_const_type sub arity = let subst_const_def sub = function | Undef inl -> Undef inl | Def c -> Def (subst_constr_subst sub c) - | OpaqueDef lc -> OpaqueDef (subst_lazy_constr sub lc) + | OpaqueDef lc -> + OpaqueDef (Future.chain ~pure:true lc (subst_lazy_constr sub)) let subst_const_body sub cb = { const_hyps = (match cb.const_hyps with [] -> [] | _ -> assert false); @@ -84,14 +85,24 @@ let hcons_const_type = function let hcons_const_def = function | Undef inl -> Undef inl - | Def cs -> Def (from_val (Term.hcons_constr (Lazyconstr.force cs))) - | OpaqueDef lc -> OpaqueDef (Lazyconstr.hcons_lazy_constr lc) + | Def l_constr -> + let constr = force l_constr in + Def (from_val (Term.hcons_constr constr)) + | OpaqueDef lc -> + if Future.is_val lc then + let constr = force_opaque (Future.force lc) in + OpaqueDef (Future.from_val (opaque_from_val (Term.hcons_constr constr))) + else OpaqueDef lc let hcons_const_body cb = { cb with const_body = hcons_const_def cb.const_body; const_type = hcons_const_type cb.const_type; - const_constraints = Univ.hcons_constraints cb.const_constraints } + const_constraints = + if Future.is_val cb.const_constraints then + Future.from_val + (Univ.hcons_constraints (Future.force cb.const_constraints)) + else cb.const_constraints } (** Inductive types *) @@ -193,3 +204,21 @@ let hcons_mind mib = mind_packets = Array.smartmap hcons_mind_packet mib.mind_packets; mind_params_ctxt = hcons_rel_context mib.mind_params_ctxt; mind_constraints = Univ.hcons_constraints mib.mind_constraints } + +let join_constant_body cb = + ignore(Future.join cb.const_constraints); + match cb.const_body with + | OpaqueDef d -> ignore(Future.join d) + | _ -> () + +let string_of_side_effect = function + | NewConstant (c,_) -> Names.string_of_con c +type side_effects = side_effect list +let no_seff = ([] : side_effects) +let iter_side_effects f l = List.iter f l +let fold_side_effects f a l = List.fold_left f a l +let uniquize_side_effects l = List.uniquize l +let union_side_effects l1 l2 = l1 @ l2 +let flatten_side_effects l = List.flatten l +let side_effects_of_list l = l +let cons_side_effects x l = x :: l diff --git a/kernel/declareops.mli b/kernel/declareops.mli index c4d67ba52..635b8c47a 100644 --- a/kernel/declareops.mli +++ b/kernel/declareops.mli @@ -28,6 +28,19 @@ val body_of_constant : constant_body -> Term.constr option val is_opaque : constant_body -> bool +(** Side effects *) + +val string_of_side_effect : side_effect -> string + +type side_effects +val no_seff : side_effects +val iter_side_effects : (side_effect -> unit) -> side_effects -> unit +val fold_side_effects : ('a -> side_effect -> 'a) -> 'a -> side_effects -> 'a +val uniquize_side_effects : side_effects -> side_effects +val union_side_effects : side_effects -> side_effects -> side_effects +val flatten_side_effects : side_effects list -> side_effects +val side_effects_of_list : side_effect list -> side_effects +val cons_side_effects : side_effect -> side_effects -> side_effects (** {6 Inductive types} *) @@ -45,6 +58,7 @@ val subst_wf_paths : substitution -> wf_paths -> wf_paths val subst_mind : substitution -> mutual_inductive_body -> mutual_inductive_body +val join_constant_body : constant_body -> unit (** {6 Hash-consing} *) diff --git a/kernel/entries.mli b/kernel/entries.mli index 392a2cd84..3b7a2fd19 100644 --- a/kernel/entries.mli +++ b/kernel/entries.mli @@ -47,7 +47,8 @@ type mutual_inductive_entry = { mind_entry_inds : one_inductive_entry list } (** {6 Constants (Definition/Axiom) } *) -type const_entry_body = constr +type proof_output = constr * Declareops.side_effects +type const_entry_body = proof_output Future.computation type definition_entry = { const_entry_body : const_entry_body; diff --git a/kernel/mod_typing.ml b/kernel/mod_typing.ml index f7f3c2b77..67db4e806 100644 --- a/kernel/mod_typing.ml +++ b/kernel/mod_typing.ml @@ -97,16 +97,17 @@ and check_with_def env sign (idl,c) mp equiv = let (j,cst1) = Typeops.infer env' c in let typ = Typeops.type_of_constant_type env' cb.const_type in let cst2 = Reduction.conv_leq env' j.uj_type typ in - let cst = - union_constraints - (union_constraints cb.const_constraints cst1) + let cst = union_constraints + (union_constraints + (Future.force cb.const_constraints) cst1) cst2 in let def = Def (Lazyconstr.from_val j.uj_val) in def,cst | Def cs -> let cst1 = Reduction.conv env' c (Lazyconstr.force cs) in - let cst = union_constraints cb.const_constraints cst1 in + let cst = union_constraints + (Future.force cb.const_constraints) cst1 in let def = Def (Lazyconstr.from_val c) in def,cst in @@ -115,7 +116,7 @@ and check_with_def env sign (idl,c) mp equiv = const_body = def; const_body_code = Cemitcodes.from_val (compile_constant_body env' def); - const_constraints = cst } + const_constraints = Future.from_val cst } in SEBstruct(before@(l,SFBconst(cb'))::after),cb',cst | _ -> @@ -373,13 +374,14 @@ let rec add_struct_expr_constraints env = function (add_struct_expr_constraints env meb1) meb2) | SEBwith(meb,With_definition_body(_,cb))-> - Environ.add_constraints cb.const_constraints + Environ.add_constraints (Future.force cb.const_constraints) (add_struct_expr_constraints env meb) | SEBwith(meb,With_module_body(_,_))-> add_struct_expr_constraints env meb and add_struct_elem_constraints env = function - | SFBconst cb -> Environ.add_constraints cb.const_constraints env + | SFBconst cb -> + Environ.add_constraints (Future.force cb.const_constraints) env | SFBmind mib -> Environ.add_constraints mib.mind_constraints env | SFBmodule mb -> add_module_constraints env mb | SFBmodtype mtb -> add_modtype_constraints env mtb @@ -418,7 +420,7 @@ let rec struct_expr_constraints cst = function meb2 | SEBwith(meb,With_definition_body(_,cb))-> struct_expr_constraints - (Univ.union_constraints cb.const_constraints cst) meb + (Univ.union_constraints (Future.force cb.const_constraints) cst) meb | SEBwith(meb,With_module_body(_,_))-> struct_expr_constraints cst meb diff --git a/kernel/modops.ml b/kernel/modops.ml index 3b789016b..39e5514c9 100644 --- a/kernel/modops.ml +++ b/kernel/modops.ml @@ -588,3 +588,32 @@ let clean_bounded_mod_expr = function let str_clean = collect_mbid [] str in if str_clean == str then str else str_clean | str -> str + +let rec join_module_body mb = + Option.iter join_struct_expr_body mb.mod_expr; + Option.iter join_struct_expr_body mb.mod_type_alg; + join_struct_expr_body mb.mod_type +and join_structure_body struc = + let join_body (l,body) = match body with + | SFBconst sb -> join_constant_body sb + | SFBmind _ -> () + | SFBmodule m -> join_module_body m + | SFBmodtype m -> + join_struct_expr_body m.typ_expr; + Option.iter join_struct_expr_body m.typ_expr_alg in + List.iter join_body struc; +and join_struct_expr_body = function + | SEBfunctor (_,t,e) -> + join_struct_expr_body t.typ_expr; + Option.iter join_struct_expr_body t.typ_expr_alg; + join_struct_expr_body e + | SEBident mp -> () + | SEBstruct s -> join_structure_body s + | SEBapply (mexpr,marg,u) -> + join_struct_expr_body mexpr; + join_struct_expr_body marg + | SEBwith (seb,wdcl) -> + join_struct_expr_body seb; + match wdcl with + | With_module_body _ -> () + | With_definition_body (_, sb) -> join_constant_body sb diff --git a/kernel/modops.mli b/kernel/modops.mli index c4d8ab349..30f9274d2 100644 --- a/kernel/modops.mli +++ b/kernel/modops.mli @@ -50,6 +50,10 @@ val subst_modtype_and_resolver : module_type_body -> module_path -> val clean_bounded_mod_expr : struct_expr_body -> struct_expr_body +val join_module_body : module_body -> unit +val join_structure_body : structure_body -> unit +val join_struct_expr_body : struct_expr_body -> unit + (** Errors *) type signature_mismatch_error = diff --git a/kernel/safe_typing.ml b/kernel/safe_typing.ml index 753b97a0c..13368aab9 100644 --- a/kernel/safe_typing.ml +++ b/kernel/safe_typing.ml @@ -99,6 +99,7 @@ type safe_environment = objlabels : Label.Set.t; revstruct : structure_body; univ : Univ.constraints; + future_cst : Univ.constraints Future.computation list; engagement : engagement option; imports : library_info list; loads : (module_path * module_body) list; @@ -107,6 +108,36 @@ type safe_environment = let exists_modlabel l senv = Label.Set.mem l senv.modlabels let exists_objlabel l senv = Label.Set.mem l senv.objlabels +(* type to be maintained isomorphic to Entries.side_effects *) +(* XXX ideally this function obtains a valid side effect that + * can be pushed into another (safe) environment without re-typechecking *) +type side_effect = NewConstant of constant * constant_body +let sideff_of_con env c = + Obj.magic (NewConstant (c, Environ.lookup_constant c env.env)) + +let env_of_safe_env senv = senv.env +let env_of_senv = env_of_safe_env + +type constraints_addition = + Now of Univ.constraints | Later of Univ.constraints Future.computation + +let add_constraints cst senv = + match cst with + | Later fc -> {senv with future_cst = fc :: senv.future_cst} + | Now cst -> + { senv with + env = Environ.add_constraints cst senv.env; + univ = Univ.union_constraints cst senv.univ } + +let is_curmod_library senv = + match senv.modinfo.variant with LIBRARY _ -> true | _ -> false + +let rec join_safe_environment e = + join_structure_body e.revstruct; + List.fold_left + (fun e fc -> add_constraints (Now (Future.join fc)) e) + {e with future_cst = []} e.future_cst + let check_modlabel l senv = if exists_modlabel l senv then error_existing_label l let check_objlabel l senv = @@ -141,25 +172,21 @@ let rec empty_environment = modlabels = Label.Set.empty; objlabels = Label.Set.empty; revstruct = []; + future_cst = []; univ = Univ.empty_constraint; engagement = None; imports = []; loads = []; local_retroknowledge = [] } -let env_of_safe_env senv = senv.env -let env_of_senv = env_of_safe_env - -let add_constraints cst senv = - { senv with - env = Environ.add_constraints cst senv.env; - univ = Univ.union_constraints cst senv.univ } - let constraints_of_sfb = function - | SFBconst cb -> cb.const_constraints - | SFBmind mib -> mib.mind_constraints - | SFBmodtype mtb -> mtb.typ_constraints - | SFBmodule mb -> mb.mod_constraints + | SFBmind mib -> Now mib.mind_constraints + | SFBmodtype mtb -> Now mtb.typ_constraints + | SFBmodule mb -> Now mb.mod_constraints + | SFBconst cb -> + match Future.peek_val cb.const_constraints with + | Some c -> Now c + | None -> Later cb.const_constraints (* A generic function for adding a new field in a same environment. It also performs the corresponding [add_constraints]. *) @@ -246,14 +273,20 @@ let safe_push_named (id,_,_ as d) env = Environ.push_named d env let push_named_def (id,de) senv = - let (c,typ,cst) = Term_typing.translate_local_def senv.env de in - let senv' = add_constraints cst senv in + let (c,typ,cst) = Term_typing.translate_local_def senv.env id de in + (* XXX for now we force *) + let c = match c with + | Def c -> Lazyconstr.force c + | OpaqueDef c -> Lazyconstr.force_opaque (Future.join c) + | _ -> assert false in + let cst = Future.join cst in + let senv' = add_constraints (Now cst) senv in let env'' = safe_push_named (id,Some c,typ) senv'.env in (cst, {senv' with env=env''}) let push_named_assum (id,t) senv = let (t,cst) = Term_typing.translate_local_assum senv.env t in - let senv' = add_constraints cst senv in + let senv' = add_constraints (Now cst) senv in let env'' = safe_push_named (id,None,t) senv'.env in (cst, {senv' with env=env''}) @@ -267,16 +300,17 @@ type global_declaration = let add_constant dir l decl senv = let kn = make_con senv.modinfo.modpath dir l in let cb = match decl with - | ConstantEntry ce -> Term_typing.translate_constant senv.env ce + | ConstantEntry ce -> Term_typing.translate_constant senv.env kn ce | GlobalRecipe r -> - let cb = Term_typing.translate_recipe senv.env r in + let cb = Term_typing.translate_recipe senv.env kn r in if DirPath.is_empty dir then Declareops.hcons_const_body cb else cb in let cb = match cb.const_body with | OpaqueDef lc when DirPath.is_empty dir -> (* In coqc, opaque constants outside sections will be stored indirectly in a specific table *) - { cb with const_body = OpaqueDef (Lazyconstr.turn_indirect lc) } + { cb with const_body = + OpaqueDef (Future.chain lc Lazyconstr.turn_indirect) } | _ -> cb in let senv' = add_field (l,SFBconst cb) (C kn) senv in @@ -317,7 +351,7 @@ let add_modtype l mte inl senv = (* full_add_module adds module with universes and constraints *) let full_add_module mb senv = - let senv = add_constraints mb.mod_constraints senv in + let senv = add_constraints (Now mb.mod_constraints) senv in { senv with env = Modops.add_module mb senv.env } (* Insertion of modules *) @@ -350,6 +384,7 @@ let start_module l senv = objlabels = Label.Set.empty; revstruct = []; univ = Univ.empty_constraint; + future_cst = []; engagement = None; imports = senv.imports; loads = []; @@ -435,6 +470,7 @@ let end_module l restype senv = objlabels = oldsenv.objlabels; revstruct = (l,SFBmodule mb)::oldsenv.revstruct; univ = Univ.union_constraints senv'.univ oldsenv.univ; + future_cst = senv'.future_cst @ oldsenv.future_cst; (* engagement is propagated to the upper level *) engagement = senv'.engagement; imports = senv'.imports; @@ -457,14 +493,14 @@ let end_module l restype senv = senv.modinfo.modpath inl me in mtb.typ_expr,mtb.typ_constraints,mtb.typ_delta in - let senv = add_constraints cst senv in + let senv = add_constraints (Now cst) senv in let mp_sup = senv.modinfo.modpath in (* Include Self support *) let rec compute_sign sign mb resolver senv = match sign with | SEBfunctor(mbid,mtb,str) -> let cst_sub = check_subtypes senv.env mb mtb in - let senv = add_constraints cst_sub senv in + let senv = add_constraints (Now cst_sub) senv in let mpsup_delta = inline_delta_resolver senv.env inl mp_sup mbid mtb mb.typ_delta in @@ -533,6 +569,7 @@ let add_module_parameter mbid mte inl senv = objlabels = senv.objlabels; revstruct = []; univ = senv.univ; + future_cst = senv.future_cst; engagement = senv.engagement; imports = senv.imports; loads = []; @@ -557,6 +594,7 @@ let start_modtype l senv = objlabels = Label.Set.empty; revstruct = []; univ = Univ.empty_constraint; + future_cst = []; engagement = None; imports = senv.imports; loads = [] ; @@ -609,6 +647,7 @@ let end_modtype l senv = objlabels = oldsenv.objlabels; revstruct = (l,SFBmodtype mtb)::oldsenv.revstruct; univ = Univ.union_constraints senv.univ oldsenv.univ; + future_cst = senv.future_cst @ oldsenv.future_cst; engagement = senv.engagement; imports = senv.imports; loads = senv.loads@oldsenv.loads; @@ -644,6 +683,8 @@ type compiled_library = { type native_library = Nativecode.global list +let join_compiled_library l = join_module_body l.comp_mod + (* We check that only initial state Require's were performed before [start_library] was called *) @@ -674,12 +715,16 @@ let start_library dir senv = objlabels = Label.Set.empty; revstruct = []; univ = Univ.empty_constraint; + future_cst = []; engagement = None; imports = senv.imports; loads = []; local_retroknowledge = [] } let export senv dir = + let senv = + try join_safe_environment senv + with e -> Errors.errorlabstrm "future" (Errors.print e) in let modinfo = senv.modinfo in begin match modinfo.variant with @@ -763,7 +808,6 @@ let import lib digest senv = in mp, senv, lib.comp_natsymbs - type judgment = unsafe_judgment let j_val j = j.uj_val @@ -787,3 +831,6 @@ let register_inline kn senv = let new_globals = { env.Pre_env.env_globals with Pre_env.env_constants = new_constants } in let env = { env with Pre_env.env_globals = new_globals } in { senv with env = env_of_pre_env env } + +let add_constraints c = add_constraints (Now c) + diff --git a/kernel/safe_typing.mli b/kernel/safe_typing.mli index 7ca033383..210d601fb 100644 --- a/kernel/safe_typing.mli +++ b/kernel/safe_typing.mli @@ -26,6 +26,15 @@ type safe_environment val env_of_safe_env : safe_environment -> Environ.env +val sideff_of_con : safe_environment -> constant -> side_effect + +val is_curmod_library : safe_environment -> bool + + +(* safe_environment has functional data affected by lazy computations, + * thus this function returns a new safe_environment *) +val join_safe_environment : safe_environment -> safe_environment + val empty_environment : safe_environment val is_empty : safe_environment -> bool @@ -109,6 +118,8 @@ val export : safe_environment -> DirPath.t val import : compiled_library -> Digest.t -> safe_environment -> module_path * safe_environment * Nativecode.symbol array +val join_compiled_library : compiled_library -> unit + (** {6 Typing judgments } *) type judgment diff --git a/kernel/term_typing.ml b/kernel/term_typing.ml index 33a4b55e8..fd31d782a 100644 --- a/kernel/term_typing.ml +++ b/kernel/term_typing.ml @@ -23,6 +23,10 @@ open Environ open Entries open Typeops +let debug = false +let prerr_endline = + if debug then prerr_endline else fun _ -> () + let constrain_type env j cst1 = function | None -> make_polymorphic_if_constant_for_ind env j, cst1 @@ -33,19 +37,6 @@ let constrain_type env j cst1 = function let cstrs = union_constraints (union_constraints cst1 cst2) cst3 in NonPolymorphicType t, cstrs -let local_constrain_type env j cst1 = function - | None -> - j.uj_type, cst1 - | Some t -> - let (tj,cst2) = infer_type env t in - let (_,cst3) = judge_of_cast env j DEFAULTcast tj in - assert (eq_constr t tj.utj_val); - t, union_constraints (union_constraints cst1 cst2) cst3 - -let translate_local_def env de = - let (j,cst) = infer env de.const_entry_body in - let (typ,cst) = local_constrain_type env j cst de.const_entry_type in - (j.uj_val,typ,cst) let translate_local_assum env t = let (j,cst) = infer env t in @@ -55,23 +46,89 @@ let translate_local_assum env t = (* Insertion of constants and parameters in environment. *) -let infer_declaration env = function +let handle_side_effects env body side_eff = + let handle_sideff t (NewConstant (c,cb)) = + let cname = + let name = string_of_con c in + for i = 0 to String.length name - 1 do + if name.[i] = '.' || name.[i] = '#' then name.[i] <- '_' done; + Name (id_of_string name) in + let rec sub i x = match kind_of_term x with + | Const c' when eq_constant c c' -> mkRel i + | _ -> map_constr_with_binders ((+) 1) sub i x in + match cb.const_body with + | Undef _ -> assert false + | Def b -> + let b = Lazyconstr.force b in + let b_ty = Typeops.type_of_constant_type env cb.const_type in + let t = sub 1 (Vars.lift 1 t) in + mkLetIn (cname, b, b_ty, t) + | OpaqueDef b -> + let b = Lazyconstr.force_opaque (Future.force b) in + let b_ty = Typeops.type_of_constant_type env cb.const_type in + let t = sub 1 (Vars.lift 1 t) in + mkApp (mkLambda (cname, b_ty, t), [|b|]) + in + (* CAVEAT: we assure a proper order *) + Declareops.fold_side_effects handle_sideff body + (Declareops.uniquize_side_effects side_eff) + +(* what is used for debugging only *) +let infer_declaration ?(what="UNKNOWN") env dcl = + match dcl with | DefinitionEntry c -> - let (j,cst) = infer env c.const_entry_body in - let j = - {uj_val = hcons_constr j.uj_val; - uj_type = hcons_constr j.uj_type} in - let (typ,cst) = constrain_type env j cst c.const_entry_type in - let def = - if c.const_entry_opaque - then OpaqueDef (Lazyconstr.opaque_from_val j.uj_val) - else Def (Lazyconstr.from_val j.uj_val) - in - def, typ, cst, c.const_entry_inline_code, c.const_entry_secctx + let ctx, entry_body = c.const_entry_secctx, c.const_entry_body in + if c.const_entry_opaque && c.const_entry_type <> None then + let _ = prerr_endline ("deferring typing of "^what) in + let body_cst = Future.chain ~id:("infer_declaration " ^ what) + entry_body (fun entry_body -> + let _ = prerr_endline ("forcing deferred typing of "^what) in + let body, side_eff = entry_body in + let _ = prerr_endline ("... got proof of "^what) in + let body = if side_eff = Declareops.no_seff then body else + let _ = prerr_endline (" Handling the following side eff:") in + Declareops.iter_side_effects (fun e -> + prerr_endline(" " ^ Declareops.string_of_side_effect e)) + side_eff; + handle_side_effects env body side_eff in + let (j,cst) = infer env body in + let _ = prerr_endline ("... typed proof of "^what) in + let j = + {uj_val = hcons_constr j.uj_val; + uj_type = hcons_constr j.uj_type} in + let (_typ,cst) = constrain_type env j cst c.const_entry_type in + Lazyconstr.opaque_from_val j.uj_val, cst) in + let body, cst = Future.split2 body_cst in + let def = OpaqueDef body in + let typ = match c.const_entry_type with + | None -> assert false + | Some typ -> NonPolymorphicType typ in + (def, typ, cst, c.const_entry_inline_code, ctx) + else + let _ = prerr_endline ("typing now "^what) in + let body, side_eff = Future.force entry_body in + let body = if side_eff = Declareops.no_seff then body else + let _ = prerr_endline (" Handling the following side eff:") in + Declareops.iter_side_effects (fun e -> + prerr_endline(" " ^ Declareops.string_of_side_effect e)) + side_eff; + handle_side_effects env body side_eff in + let (j,cst) = + try infer env body + with Not_found as e -> + prerr_endline ("infer not found " ^ what); + raise e in + let j = + {uj_val = hcons_constr j.uj_val; + uj_type = hcons_constr j.uj_type} in + let (typ,cst) = constrain_type env j cst c.const_entry_type in + let _ = prerr_endline ("...typed "^what) in + let def = Def (Lazyconstr.from_val j.uj_val) in + (def, typ, Future.from_val cst, c.const_entry_inline_code, ctx) | ParameterEntry (ctx,t,nl) -> let (j,cst) = infer env t in let t = hcons_constr (Typeops.assumption_of_judgment env j) in - Undef nl, NonPolymorphicType t, cst, false, ctx + Undef nl, NonPolymorphicType t, Future.from_val cst, false, ctx let global_vars_set_constant_type env = function | NonPolymorphicType t -> global_vars_set env t @@ -81,48 +138,78 @@ let global_vars_set_constant_type env = function (fun t c -> Id.Set.union (global_vars_set env t) c)) ctx ~init:Id.Set.empty -let check_declared_variables declared inferred = - let mk_set l = List.fold_right Id.Set.add (List.map pi1 l) Id.Set.empty in - let undeclared_set = Id.Set.diff (mk_set inferred) (mk_set declared) in - if not (Id.Set.is_empty undeclared_set) then - error ("The following section variables are used but not declared:\n"^ - (String.concat ", " - (List.map Id.to_string (Id.Set.elements undeclared_set)))) - -let build_constant_declaration env (def,typ,cst,inline_code,ctx) = - let hyps = - let inferred = - let ids_typ = global_vars_set_constant_type env typ in - let ids_def = match def with - | Undef _ -> Id.Set.empty - | Def cs -> global_vars_set env (Lazyconstr.force cs) - | OpaqueDef lc -> global_vars_set env (Lazyconstr.force_opaque lc) - in - keep_hyps env (Id.Set.union ids_typ ids_def) - in +let build_constant_declaration kn env (def,typ,cst,inline_code,ctx) = + let check declared inferred = + let mk_set l = List.fold_right Id.Set.add (List.map pi1 l) Id.Set.empty in + let inferred_set, declared_set = mk_set inferred, mk_set declared in + if not (Id.Set.subset inferred_set declared_set) then + error ("The following section variable are used but not declared:\n"^ + (String.concat ", " (List.map Id.to_string + (Id.Set.elements (Idset.diff inferred_set declared_set))))) in + (* We try to postpone the computation of used section variables *) + let hyps, def = match ctx with - | None -> inferred - | Some declared -> - check_declared_variables declared inferred; - declared - in - let tps = Cemitcodes.from_val (compile_constant_body env def) in + | None when named_context env <> [] -> + (* No declared section vars, and non-empty section context: + we must look at the body NOW, if any *) + let ids_typ = global_vars_set_constant_type env typ in + let ids_def = match def with + | Undef _ -> Idset.empty + | Def cs -> global_vars_set env (Lazyconstr.force cs) + | OpaqueDef lc -> + (* we force so that cst are added to the env immediately after *) + ignore(Future.join cst); + global_vars_set env (Lazyconstr.force_opaque (Future.force lc)) in + keep_hyps env (Idset.union ids_typ ids_def), def + | None -> [], def (* Empty section context: no need to check *) + | Some declared -> + (* We use the declared set and chain a check of correctness *) + declared, + match def with + | Undef _ as x -> x (* nothing to check *) + | Def cs as x -> + let ids_typ = global_vars_set_constant_type env typ in + let ids_def = global_vars_set env (Lazyconstr.force cs) in + let inferred = keep_hyps env (Idset.union ids_typ ids_def) in + check declared inferred; + x + | OpaqueDef lc -> (* In this case we can postpone the check *) + OpaqueDef (Future.chain ~id:(string_of_con kn) lc (fun lc -> + let ids_typ = global_vars_set_constant_type env typ in + let ids_def = + global_vars_set env (Lazyconstr.force_opaque lc) in + let inferred = keep_hyps env (Idset.union ids_typ ids_def) in + check declared inferred; lc)) in { const_hyps = hyps; const_body = def; const_type = typ; - const_body_code = tps; + const_body_code = Cemitcodes.from_val (compile_constant_body env def); const_constraints = cst; const_native_name = ref NotLinked; const_inline_code = inline_code } (*s Global and local constant declaration. *) -let translate_constant env ce = - build_constant_declaration env (infer_declaration env ce) +let translate_constant env kn ce = + build_constant_declaration kn env + (infer_declaration ~what:(string_of_con kn) env ce) + +let translate_recipe env kn r = + let def,typ,cst,inline_code,hyps = Cooking.cook_constant env r in + build_constant_declaration kn env (def,typ,cst,inline_code,hyps) -let translate_recipe env r = - build_constant_declaration env (Cooking.cook_constant env r) +let translate_local_def env id centry = + let def,typ,cst,inline_code,ctx = + infer_declaration ~what:(string_of_id id) env (DefinitionEntry centry) in + let typ = type_of_constant_type env typ in + def, typ, cst (* Insertion of inductive types. *) let translate_mind env kn mie = Indtypes.check_inductive env kn mie + +let handle_side_effects env ce = { ce with + const_entry_body = Future.chain ~id:"handle_side_effects" + ce.const_entry_body (fun (body, side_eff) -> + handle_side_effects env body side_eff, Declareops.no_seff); +} diff --git a/kernel/term_typing.mli b/kernel/term_typing.mli index c9bff84fc..59706bb83 100644 --- a/kernel/term_typing.mli +++ b/kernel/term_typing.mli @@ -13,20 +13,25 @@ open Environ open Declarations open Entries -val translate_local_def : env -> definition_entry -> - constr * types * constraints +val translate_local_def : env -> Id.t -> definition_entry -> + constant_def * types * constant_constraints val translate_local_assum : env -> types -> types * constraints -val translate_constant : env -> constant_entry -> constant_body +(* returns the same definition_entry but with side effects turned into + * let-ins or beta redexes. it is meant to get a term out of a not yet + * type checked proof *) +val handle_side_effects : env -> definition_entry -> definition_entry + +val translate_constant : env -> constant -> constant_entry -> constant_body val translate_mind : env -> mutual_inductive -> mutual_inductive_entry -> mutual_inductive_body -val translate_recipe : env -> Cooking.recipe -> constant_body - +val translate_recipe : env -> constant -> Cooking.recipe -> constant_body (** Internal functions, mentioned here for debug purpose only *) -val infer_declaration : env -> constant_entry -> Cooking.result -val build_constant_declaration : env -> Cooking.result -> constant_body +val infer_declaration : ?what:string -> env -> constant_entry -> Cooking.result +val build_constant_declaration : + constant -> env -> Cooking.result -> constant_body diff --git a/lib/flags.ml b/lib/flags.ml index 06c52e4b2..d764894fa 100644 --- a/lib/flags.ml +++ b/lib/flags.ml @@ -55,6 +55,8 @@ let xml_export = ref false let ide_slave = ref false +let time = ref false + type load_proofs = Force | Lazy | Dont let load_proofs = ref Lazy diff --git a/lib/flags.mli b/lib/flags.mli index 14b9f26bc..388b33c12 100644 --- a/lib/flags.mli +++ b/lib/flags.mli @@ -22,6 +22,8 @@ val xml_export : bool ref val ide_slave : bool ref +val time : bool ref + type load_proofs = Force | Lazy | Dont val load_proofs : load_proofs ref diff --git a/lib/vcs.ml b/lib/vcs.ml index 2202ddeb0..dda0384f0 100644 --- a/lib/vcs.ml +++ b/lib/vcs.ml @@ -28,7 +28,7 @@ module type S = sig type ('kind,'diff,'info) t constraint 'kind = [> `Master ] - val empty : default_info:'info -> id -> ('kind,'diff,'info) t + val empty : id -> ('kind,'diff,'info) t val current_branch : ('k,'e,'i) t -> branch_name val branches : ('k,'e,'i) t -> branch_name list @@ -45,7 +45,7 @@ module type S = sig val checkout : ('k,'e,'i) t -> branch_name -> ('k,'e,'i) t val set_info : ('k,'e,'info) t -> id -> 'info -> ('k,'e,'info) t - val get_info : ('k,'e,'info) t -> id -> 'info + val get_info : ('k,'e,'info) t -> id -> 'info option val create_cluster : ('k,'e,'i) t -> id list -> ('k,'e,'i) t @@ -78,14 +78,12 @@ type ('kind,'edge,'info) t = { cur_branch : branch_name; heads : (branch_name * 'kind branch_info) list; dag : ('edge,'info) Dag.t; - default_info : 'info; } -let empty ~default_info root = { +let empty root = { cur_branch = master; heads = [ master, { root = root; pos = root; kind = `Master } ]; dag = Dag.empty; - default_info; } let add_node vcs id edges = @@ -113,12 +111,10 @@ let delete_branch vcs name = { vcs with let merge vcs id ~ours:tr1 ~theirs:tr2 ?(into = vcs.cur_branch) name = assert (name <> into); - let local = into = vcs.cur_branch in let br_id = (get_branch vcs name).pos in let cur_id = (get_branch vcs into).pos in let vcs = add_node vcs id [tr1,cur_id; tr2,br_id] in let vcs = reset_branch vcs into id in - let vcs = if local then reset_branch vcs name id else vcs in vcs let commit vcs id tr = @@ -129,8 +125,7 @@ let commit vcs id tr = let checkout vcs name = { vcs with cur_branch = name } let set_info vcs id info = { vcs with dag = Dag.set_info vcs.dag id info } -let get_info vcs id = - match Dag.get_info vcs.dag id with Some x -> x | None -> vcs.default_info +let get_info vcs id = Dag.get_info vcs.dag id let create_cluster vcs l = { vcs with dag = Dag.create_cluster vcs.dag l } diff --git a/lib/vcs.mli b/lib/vcs.mli index 7c82352ae..7ac568c12 100644 --- a/lib/vcs.mli +++ b/lib/vcs.mli @@ -43,7 +43,7 @@ module type S = sig type ('kind,'diff,'info) t constraint 'kind = [> `Master ] - val empty : default_info:'info -> id -> ('kind,'diff,'info) t + val empty : id -> ('kind,'diff,'info) t val current_branch : ('k,'e,'i) t -> branch_name val branches : ('k,'e,'i) t -> branch_name list @@ -60,7 +60,7 @@ module type S = sig val checkout : ('k,'e,'i) t -> branch_name -> ('k,'e,'i) t val set_info : ('k,'e,'info) t -> id -> 'info -> ('k,'e,'info) t - val get_info : ('k,'e,'info) t -> id -> 'info + val get_info : ('k,'e,'info) t -> id -> 'info option val create_cluster : ('k,'e,'i) t -> id list -> ('k,'e,'i) t diff --git a/library/global.ml b/library/global.ml index 28b691769..5425e5930 100644 --- a/library/global.ml +++ b/library/global.ml @@ -14,7 +14,17 @@ open Safe_typing (* We introduce here the global environment of the system, and we declare it as a synchronized table. *) -let global_env = Summary.ref empty_environment ~name:"Global environment" +let global_env = ref empty_environment + +let join_safe_environment () = + global_env := Safe_typing.join_safe_environment !global_env + +let () = + Summary.declare_summary "Global environment" + { Summary.freeze_function = (fun b -> + if b then join_safe_environment (); !global_env); + unfreeze_function = (fun fr -> global_env := fr); + init_function = (fun () -> global_env := empty_environment) } let safe_env () = !global_env diff --git a/library/global.mli b/library/global.mli index 3238294f2..a5ca92013 100644 --- a/library/global.mli +++ b/library/global.mli @@ -102,6 +102,8 @@ val import : compiled_library -> Digest.t -> val type_of_global : Globnames.global_reference -> types val env_of_context : Environ.named_context_val -> Environ.env +val join_safe_environment : unit -> unit + (** spiwack: register/unregister function for retroknowledge *) val register : Retroknowledge.field -> constr -> constr -> unit diff --git a/library/lib.ml b/library/lib.ml index 47341e675..b4371812b 100644 --- a/library/lib.ml +++ b/library/lib.ml @@ -523,85 +523,8 @@ let close_section () = 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 (); - (* Always at least one frozen state in the libstack *) - match find_entry_p is_frozen_state with - | (sp, FrozenState f) -> - Summary.unfreeze_summaries f; - let (after,_,_) = split_lib sp in - recache_context after - | _ -> assert false - -let reset_to test = - let (_,_,before) = split_lib_gen test in - set_lib_stk before - -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 String.equal (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 String.equal (object_tag o) "DOT" && - Int.equal 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 (is_label_n n); - (* forget state numbers after n only if reset succeeded *) - reset_command_label n - -let raw_reset_initial () = reset_label first_command_label - -(** 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 && String.equal (object_tag o) "DOT" -> true - | ((fp, _),_) -> - let (_, name) = repr_path fp in - let () = if Names.Id.equal id name then found := true in - false - in - match find_entry_p search with - | (_,Leaf o) -> outLabel o - | _ -> raise Not_found - (* State and initialization. *) type frozen = Names.DirPath.t option * library_segment @@ -616,8 +539,7 @@ let unfreeze (mn,stk) = let init () = unfreeze (None,[]); Summary.init_summaries (); - add_frozen_state (); (* Stores e.g. the keywords declared in g_*.ml4 *) - mark_end_of_command () (* For allowing Reset Initial in coqtop -l *) + add_frozen_state () (* Stores e.g. the keywords declared in g_*.ml4 *) (* Misc *) diff --git a/library/lib.mli b/library/lib.mli index a956ff5d0..04f8d0775 100644 --- a/library/lib.mli +++ b/library/lib.mli @@ -147,31 +147,6 @@ val remove_section_part : Globnames.global_reference -> Names.DirPath.t 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 - -(** [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 - -(** [raw_reset_initial] is now [reset_label] to the first label. - This is meant to be used during initial Load's and compilations. - Otherwise, consider instead [Backtrack.reset_initial] *) -val raw_reset_initial : unit -> unit - -(** search the label registered immediately before adding some definition *) -val label_before_name : Names.Id.t Loc.located -> int - (** {6 We can get and set the state of the operations (used in [States]). } *) type frozen diff --git a/library/library.ml b/library/library.ml index 472b1fb1d..0766a62d7 100644 --- a/library/library.ml +++ b/library/library.ml @@ -44,17 +44,20 @@ module LibraryOrdered = DirPath module LibraryMap = Map.Make(LibraryOrdered) module LibraryFilenameMap = Map.Make(LibraryOrdered) +let join x = Safe_typing.join_compiled_library x.library_compiled (* This is a map from names to loaded libraries *) -let libraries_table = Summary.ref LibraryMap.empty ~name:"LIBRARY" +let freeze b l = if b then (LibraryMap.iter (fun _ x -> join x) l; l) else l +let libraries_table = Summary.ref LibraryMap.empty ~name:"LIBRARY" ~freeze (* 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 = Summary.ref [] ~name:"LIBRARY-LOAD" -let libraries_imports_list = Summary.ref [] ~name:"LIBRARY-IMPORT" -let libraries_exports_list = Summary.ref [] ~name:"LIBRARY-EXPORT" +let freeze b l = if b then (List.iter join l; l) else l +let libraries_loaded_list = Summary.ref [] ~name:"LIBRARY-LOAD" ~freeze +let libraries_imports_list = Summary.ref [] ~name:"LIBRARY-IMPORT" ~freeze +let libraries_exports_list = Summary.ref [] ~name:"LIBRARY-EXPORT" ~freeze (* various requests to the tables *) diff --git a/library/nametab.ml b/library/nametab.ml index 9f2bede75..fd27ef264 100644 --- a/library/nametab.ml +++ b/library/nametab.ml @@ -522,8 +522,8 @@ 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 diff --git a/library/states.ml b/library/states.ml index 1fd3fa2e5..93b2c120e 100644 --- a/library/states.ml +++ b/library/states.ml @@ -32,17 +32,6 @@ let (extern_state,intern_state) = (* Rollback. *) -let with_heavy_rollback f h x = - let st = freeze ~marshallable:false in - try - f x - with reraise -> - let e = h reraise in (unfreeze st; raise e) - -let without_rollback f h x = - try f x - with reraise -> raise (h reraise) - let with_state_protection f x = let st = freeze ~marshallable:false in try diff --git a/library/states.mli b/library/states.mli index d186bf1c8..da6f33d65 100644 --- a/library/states.mli +++ b/library/states.mli @@ -22,18 +22,6 @@ val unfreeze : state -> unit (** {6 Rollback } *) -(** [with_heavy_rollback f h x] applies [f] to [x]. If this application - ends on an exception, the wrapper [h] is applied to it, then - the state of the whole system is restored as it was before applying [f], - and finally the exception produced by [h] is raised. *) - -val with_heavy_rollback : ('a -> 'b) -> (exn -> exn) -> 'a -> 'b - -(** [without_rollback] is just like [with_heavy_rollback], except that - no state is restored in case of exception. *) - -val without_rollback : ('a -> 'b) -> (exn -> exn) -> 'a -> 'b - (** [with_state_protection f x] applies [f] to [x] and restores the state of the whole system as it was before applying [f] *) diff --git a/parsing/g_vernac.ml4 b/parsing/g_vernac.ml4 index f4380c9d7..97b0ab51c 100644 --- a/parsing/g_vernac.ml4 +++ b/parsing/g_vernac.ml4 @@ -76,6 +76,14 @@ GEXTEND Gram | IDENT "Fail"; v = vernac -> VernacFail v | IDENT "Local"; v = vernac_aux -> VernacLocal (true, v) | IDENT "Global"; v = vernac_aux -> VernacLocal (false, v) + + (* Stm backdoor *) + | IDENT "Stm"; IDENT "JoinDocument"; "." -> VernacStm JoinDocument + | IDENT "Stm"; IDENT "Finish"; "." -> VernacStm Finish + | IDENT "Stm"; IDENT "Observe"; id = INT; "." -> + VernacStm (Observe (Stateid.state_id_of_int (int_of_string id))) + | IDENT "Stm"; IDENT "Command"; v = vernac_aux -> VernacStm (Command v) + | v = vernac_aux -> v ] ] ; diff --git a/plugins/decl_mode/decl_mode.ml b/plugins/decl_mode/decl_mode.ml index 8c5aec38b..0cbd26316 100644 --- a/plugins/decl_mode/decl_mode.ml +++ b/plugins/decl_mode/decl_mode.ml @@ -99,11 +99,13 @@ let proof_cond = Proof.no_cond proof_focus let focus p = let inf = get_stack p in - Proof.focus proof_cond inf 1 p + Proof_global.with_current_proof (fun _ -> Proof.focus proof_cond inf 1) -let unfocus = Proof.unfocus proof_focus +let unfocus () = + Proof_global.with_current_proof (fun _ p -> Proof.unfocus proof_focus p ()) -let maximal_unfocus = Proof_global.maximal_unfocus proof_focus +let maximal_unfocus () = + Proof_global.with_current_proof (fun _ -> Proof.maximal_unfocus proof_focus) let get_top_stack pts = try diff --git a/plugins/decl_mode/decl_mode.mli b/plugins/decl_mode/decl_mode.mli index c54dcdc28..8e691f508 100644 --- a/plugins/decl_mode/decl_mode.mli +++ b/plugins/decl_mode/decl_mode.mli @@ -75,6 +75,6 @@ val get_last: Environ.env -> Id.t val focus : Proof.proof -> unit -val unfocus : Proof.proof -> unit +val unfocus : unit -> unit -val maximal_unfocus : Proof.proof -> unit +val maximal_unfocus : unit -> unit diff --git a/plugins/decl_mode/decl_proof_instr.ml b/plugins/decl_mode/decl_proof_instr.ml index a06b44083..e3ef0671c 100644 --- a/plugins/decl_mode/decl_proof_instr.ml +++ b/plugins/decl_mode/decl_proof_instr.ml @@ -51,12 +51,13 @@ let _ = let tcl_change_info_gen info_gen = (fun gls -> + let it, eff = sig_it gls, sig_eff gls in let concl = pf_concl gls in - let hyps = Goal.V82.hyps (project gls) (sig_it gls) in - let extra = Goal.V82.extra (project gls) (sig_it gls) in + let hyps = Goal.V82.hyps (project gls) it in + let extra = Goal.V82.extra (project gls) it in let (gl,ev,sigma) = Goal.V82.mk_goal (project gls) hyps concl (info_gen extra) in - let sigma = Goal.V82.partial_solution sigma (sig_it gls) ev in - { it = [gl] ; sigma= sigma } ) + let sigma = Goal.V82.partial_solution sigma it ev in + { it = [gl] ; sigma= sigma; eff = eff } ) let tcl_change_info info gls = let info_gen s = Store.set s Decl_mode.info info in @@ -128,34 +129,34 @@ let go_to_proof_mode () = let daimon_tac gls = set_daimon_flag (); - {it=[];sigma=sig_sig gls} + {it=[];sigma=sig_sig gls;eff=gls.eff} (* post-instruction focus management *) (* spiwack: This used to fail if there was no focusing command above, but I don't think it ever happened. I hope it doesn't mess things up*) -let goto_current_focus pts = - Decl_mode.maximal_unfocus pts +let goto_current_focus () = + Decl_mode.maximal_unfocus () -let goto_current_focus_or_top pts = - goto_current_focus pts +let goto_current_focus_or_top () = + goto_current_focus () (* return *) -let close_tactic_mode pts = - try goto_current_focus pts +let close_tactic_mode () = + try goto_current_focus () with Not_found -> error "\"return\" cannot be used outside of Declarative Proof Mode." let return_from_tactic_mode () = - close_tactic_mode (Proof_global.give_me_the_proof ()) + close_tactic_mode () (* end proof/claim *) let close_block bt pts = if Proof.no_focused_goal pts then - goto_current_focus pts + goto_current_focus () else let stack = if Proof.is_done pts then @@ -165,7 +166,7 @@ let close_block bt pts = in match bt,stack with B_claim, Claim::_ | B_focus, Focus_claim::_ | B_proof, [] -> - (goto_current_focus pts) + (goto_current_focus ()) | _, Claim::_ -> error "\"end claim\" expected." | _, Focus_claim::_ -> @@ -190,13 +191,13 @@ let close_previous_case pts = match get_top_stack pts with Per (et,_,_,_) :: _ -> anomaly (Pp.str "Weird case occured ...") | Suppose_case :: Per (et,_,_,_) :: _ -> - goto_current_focus (pts) + goto_current_focus () | _ -> error "Not inside a proof per cases or induction." else match get_stack pts with Per (et,_,_,_) :: _ -> () | Suppose_case :: Per (et,_,_,_) :: _ -> - goto_current_focus ((pts)) + goto_current_focus () | _ -> error "Not inside a proof per cases or induction." (* Proof instructions *) @@ -1443,21 +1444,21 @@ let rec postprocess pts instr = in try Inductiveops.control_only_guard env pfterm; - goto_current_focus_or_top pts + goto_current_focus_or_top () with Type_errors.TypeError(env, Type_errors.IllFormedRecBody(_,_,_,_,_)) -> anomaly (Pp.str "\"end induction\" generated an ill-formed fixpoint") end | Pend _ -> - goto_current_focus_or_top (pts) + goto_current_focus_or_top () let do_instr raw_instr pts = let has_tactic = preprocess pts raw_instr.instr in begin if has_tactic then - let { it=gls ; sigma=sigma } = Proof.V82.subgoals pts in - let gl = { it=List.hd gls ; sigma=sigma } in + let { it=gls ; sigma=sigma; eff=eff } = Proof.V82.subgoals pts in + let gl = { it=List.hd gls ; sigma=sigma; eff=eff } in let env= pf_env gl in let ist = {ltacvars = Id.Set.empty; ltacrecvars = Id.Map.empty; gsigma = sigma; genv = env} in @@ -1469,7 +1470,7 @@ let do_instr raw_instr pts = postprocess pts raw_instr.instr; (* spiwack: this should restore a compatible semantics with v8.3 where we never stayed focused on 0 goal. *) - Decl_mode.maximal_unfocus pts + Decl_mode.maximal_unfocus () let proof_instr raw_instr = let p = Proof_global.give_me_the_proof () in diff --git a/plugins/decl_mode/g_decl_mode.ml4 b/plugins/decl_mode/g_decl_mode.ml4 index 8f570af7e..072737dab 100644 --- a/plugins/decl_mode/g_decl_mode.ml4 +++ b/plugins/decl_mode/g_decl_mode.ml4 @@ -65,7 +65,7 @@ let vernac_decl_proof () = if Proof.is_done pf then Errors.error "Nothing left to prove here." else - Proof.transaction pf begin fun () -> + begin Decl_proof_instr.go_to_proof_mode () ; Proof_global.set_proof_mode "Declarative" ; Vernacentries.print_subgoals () @@ -73,14 +73,14 @@ let vernac_decl_proof () = (* spiwack: some bureaucracy is not performed here *) let vernac_return () = - Proof.transaction (Proof_global.give_me_the_proof ()) begin fun () -> + begin Decl_proof_instr.return_from_tactic_mode () ; Proof_global.set_proof_mode "Declarative" ; Vernacentries.print_subgoals () end let vernac_proof_instr instr = - Proof.transaction (Proof_global.give_me_the_proof ()) begin fun () -> + begin Decl_proof_instr.proof_instr instr; Vernacentries.print_subgoals () end @@ -397,4 +397,10 @@ GLOBAL: proof_instr; ; END;; - +open Vernacexpr + +let () = + Vernac_classifier.declare_vernac_classifier "decl_mode" (function + | VernacExtend (("DeclProof"|"DeclReturn"), _) -> VtProofStep, VtNow + | VernacExtend ("ProofInstr", _) -> VtProofStep, VtLater + | _ -> VtUnknown, VtNow) diff --git a/plugins/extraction/extraction.ml b/plugins/extraction/extraction.ml index 5c67c33b6..68f832997 100644 --- a/plugins/extraction/extraction.ml +++ b/plugins/extraction/extraction.ml @@ -993,7 +993,7 @@ let extract_constant env kn cb = | OpaqueDef c -> add_opaque r; if access_opaque () then - mk_typ (Lazyconstr.force_opaque c) + mk_typ (Lazyconstr.force_opaque (Future.force c)) else mk_typ_ax ()) | (Info,Default) -> (match cb.const_body with @@ -1002,7 +1002,7 @@ let extract_constant env kn cb = | OpaqueDef c -> add_opaque r; if access_opaque () then - mk_def (Lazyconstr.force_opaque c) + mk_def (Lazyconstr.force_opaque (Future.force c)) else mk_ax ()) let extract_constant_spec env kn cb = diff --git a/plugins/funind/functional_principles_types.ml b/plugins/funind/functional_principles_types.ml index f70ce0092..efed9a856 100644 --- a/plugins/funind/functional_principles_types.ml +++ b/plugins/funind/functional_principles_types.ml @@ -339,7 +339,8 @@ let generate_functional_principle let value = change_property_sort s new_principle_type new_princ_name in (* Pp.msgnl (str "new principle := " ++ pr_lconstr value); *) let ce = - { const_entry_body = value; + { const_entry_body = + Future.from_val (value,Declareops.no_seff); const_entry_secctx = None; const_entry_type = None; const_entry_opaque = false; @@ -556,7 +557,7 @@ let make_scheme (fas : (constant*glob_sort) list) : Entries.definition_entry lis List.map (compute_new_princ_type_from_rel funs sorts) other_princ_types in let first_princ_body,first_princ_type = const.Entries.const_entry_body, const.Entries.const_entry_type in - let ctxt,fix = decompose_lam_assum first_princ_body in (* the principle has for forall ...., fix .*) + let ctxt,fix = decompose_lam_assum (fst(Future.force first_princ_body)) in (* the principle has for forall ...., fix .*) let (idxs,_),(_,ta,_ as decl) = destFix fix in let other_result = List.map (* we can now compute the other principles *) @@ -598,7 +599,8 @@ let make_scheme (fas : (constant*glob_sort) list) : Entries.definition_entry lis Termops.it_mkLambda_or_LetIn (mkFix((idxs,i),decl)) ctxt in {const with - Entries.const_entry_body = princ_body; + Entries.const_entry_body = + (Future.from_val (princ_body,Declareops.no_seff)); Entries.const_entry_type = Some scheme_type } ) diff --git a/plugins/funind/invfun.ml b/plugins/funind/invfun.ml index fd074386e..7d14d1408 100644 --- a/plugins/funind/invfun.ml +++ b/plugins/funind/invfun.ml @@ -1013,10 +1013,9 @@ let do_save () = Lemmas.save_named false *) let derive_correctness make_scheme functional_induction (funs: constant list) (graphs:inductive list) = - let previous_state = States.freeze ~marshallable:false in let funs = Array.of_list funs and graphs = Array.of_list graphs in let funs_constr = Array.map mkConst funs in - try + States.with_state_protection (fun () -> let graphs_constr = Array.map mkInd graphs in let lemmas_types_infos = Util.Array.map2_i @@ -1044,7 +1043,7 @@ let derive_correctness make_scheme functional_induction (funs: constant list) (g Array.of_list (List.map (fun entry -> - (entry.Entries.const_entry_body, Option.get entry.Entries.const_entry_type ) + (fst(Future.force entry.Entries.const_entry_body), Option.get entry.Entries.const_entry_type ) ) (make_scheme (Array.map_to_list (fun const -> const,GType None) funs)) ) @@ -1122,16 +1121,8 @@ let derive_correctness make_scheme functional_induction (funs: constant list) (g let lem_cst = destConst (Constrintern.global_reference lem_id) in update_Function {finfo with completeness_lemma = Some lem_cst} ) - funs; - with reraise -> - (* In case of problem, we reset all the lemmas *) - Pfedit.delete_all_proofs (); - States.unfreeze previous_state; - raise reraise - - - - + funs) + () (***********************************************) diff --git a/plugins/funind/recdef.ml b/plugins/funind/recdef.ml index 4b9704c2c..68b291ff9 100644 --- a/plugins/funind/recdef.ml +++ b/plugins/funind/recdef.ml @@ -57,7 +57,8 @@ let find_reference sl s = let (declare_fun : Id.t -> logical_kind -> constr -> global_reference) = fun f_id kind value -> - let ce = {const_entry_body = value; + let ce = {const_entry_body = Future.from_val + (value, Declareops.no_seff); const_entry_secctx = None; const_entry_type = None; const_entry_opaque = false; @@ -1261,7 +1262,7 @@ let open_new_goal (build_proof:tactic -> tactic -> unit) using_lemmas ref_ goal_ ref_ := Some lemma ; let lid = ref [] in let h_num = ref (-1) in - ignore (Flags.silently Vernacentries.interp (Vernacexpr.VernacAbort None)); + ignore (Flags.silently Vernacentries.interp (Loc.ghost,Vernacexpr.VernacAbort None)); build_proof ( fun gls -> let hid = next_ident_away_in_goal h_id (pf_ids_of_hyps gls) in @@ -1443,7 +1444,6 @@ let (com_eqn : int -> Id.t -> let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num eq generate_induction_principle using_lemmas : unit = - let previous_label = Lib.current_command_label () in let function_type = interp_constr Evd.empty (Global.env()) type_of_f in let env = push_named (function_name,None,function_type) (Global.env()) in (* Pp.msgnl (str "function type := " ++ Printer.pr_lconstr function_type); *) @@ -1513,7 +1513,7 @@ let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num spc () ++ str"is defined" ) ) in - try + States.with_state_protection (fun () -> com_terminate tcc_lemma_name tcc_lemma_constr @@ -1523,9 +1523,6 @@ let recursive_definition is_mes function_name rec_impls type_of_f r rec_arg_num term_id using_lemmas (List.length res_vars) - hook - with reraise -> - ignore (Backtrack.backto previous_label); - (* anomaly (Pp.str "Cannot create termination Lemma") *) - raise reraise + hook) + () diff --git a/plugins/setoid_ring/newring.ml4 b/plugins/setoid_ring/newring.ml4 index aecbbfe85..be661587a 100644 --- a/plugins/setoid_ring/newring.ml4 +++ b/plugins/setoid_ring/newring.ml4 @@ -176,7 +176,7 @@ let carg c = TacDynamic(Loc.ghost,Pretyping.constr_in c) let dummy_goal env = let (gl,_,sigma) = Goal.V82.mk_goal Evd.empty (named_context_val env) mkProp Evd.Store.empty in - {Evd.it = gl; + {Evd.it = gl; Evd.eff=Declareops.no_seff; Evd.sigma = sigma} let exec_tactic env n f args = @@ -679,8 +679,10 @@ let add_theory name rth eqth morphth cst_tac (pre,post) power sign div = let lemma1 = constr_of params.(3) in let lemma2 = constr_of params.(4) in - let lemma1 = decl_constant (Id.to_string name^"_ring_lemma1") lemma1 in - let lemma2 = decl_constant (Id.to_string name^"_ring_lemma2") lemma2 in + let lemma1 = + decl_constant (Id.to_string name^"_ring_lemma1") (Future.from_val ( lemma1,Declareops.no_seff)) in + let lemma2 = + decl_constant (Id.to_string name^"_ring_lemma2") (Future.from_val ( lemma2,Declareops.no_seff)) in let cst_tac = interp_cst_tac env sigma morphth kind (zero,one,add,mul,opp) cst_tac in let pretac = @@ -1030,11 +1032,11 @@ let add_field_theory name fth eqth morphth cst_tac inj (pre,post) power sign odi match inj with | Some thm -> mkApp(constr_of params.(8),[|thm|]) | None -> constr_of params.(7) in - let lemma1 = decl_constant (Id.to_string name^"_field_lemma1") lemma1 in - let lemma2 = decl_constant (Id.to_string name^"_field_lemma2") lemma2 in - let lemma3 = decl_constant (Id.to_string name^"_field_lemma3") lemma3 in - let lemma4 = decl_constant (Id.to_string name^"_field_lemma4") lemma4 in - let cond_lemma = decl_constant (Id.to_string name^"_lemma5") cond_lemma in + let lemma1 = decl_constant (Id.to_string name^"_field_lemma1") (Future.from_val (lemma1,Declareops.no_seff)) in + let lemma2 = decl_constant (Id.to_string name^"_field_lemma2") (Future.from_val (lemma2,Declareops.no_seff)) in + let lemma3 = decl_constant (Id.to_string name^"_field_lemma3") (Future.from_val (lemma3,Declareops.no_seff)) in + let lemma4 = decl_constant (Id.to_string name^"_field_lemma4") (Future.from_val (lemma4,Declareops.no_seff)) in + let cond_lemma = decl_constant (Id.to_string name^"_lemma5") (Future.from_val (cond_lemma,Declareops.no_seff)) in let cst_tac = interp_cst_tac env sigma morphth kind (zero,one,add,mul,opp) cst_tac in let pretac = diff --git a/pretyping/evd.ml b/pretyping/evd.ml index 50d8b0c8c..3b4aa094f 100644 --- a/pretyping/evd.ml +++ b/pretyping/evd.ml @@ -683,10 +683,15 @@ type open_constr = evar_map * constr type ['a] *) type 'a sigma = { it : 'a ; - sigma : evar_map} + eff : Declareops.side_effects; + sigma : evar_map +} let sig_it x = x.it +let sig_eff x = x.eff let sig_sig x = x.sigma +let emit_side_effects eff x = + { x with eff = Declareops.union_side_effects eff x.eff } (**********************************************************) (* Failure explanation *) diff --git a/pretyping/evd.mli b/pretyping/evd.mli index f29a676c6..e67b1f81b 100644 --- a/pretyping/evd.mli +++ b/pretyping/evd.mli @@ -246,10 +246,14 @@ type open_constr = evar_map * constr type ['a] *) type 'a sigma = { it : 'a ; - sigma : evar_map} + eff : Declareops.side_effects; + sigma : evar_map +} val sig_it : 'a sigma -> 'a +val sig_eff : 'a sigma -> Declareops.side_effects val sig_sig : 'a sigma -> evar_map +val emit_side_effects : Declareops.side_effects -> 'a sigma -> 'a sigma (********************************************************* Failure explanation *) diff --git a/printing/ppvernac.ml b/printing/ppvernac.ml index 2ba972b5d..eae6d09d6 100644 --- a/printing/ppvernac.ml +++ b/printing/ppvernac.ml @@ -389,7 +389,6 @@ let pr_priority = function (* Pretty printer for vernac commands *) (**************************************) let make_pr_vernac pr_constr pr_lconstr = - let pr_constrarg c = spc () ++ pr_constr c in let pr_lconstrarg c = spc () ++ pr_lconstr c in let pr_intarg n = spc () ++ int n in @@ -477,6 +476,13 @@ let rec pr_vernac = function | VernacProgram v -> str"Program" ++ spc() ++ pr_vernac v | VernacLocal (local, v) -> pr_locality local ++ spc() ++ pr_vernac v + (* Stm *) + | VernacStm JoinDocument -> str"Stm JoinDocument" + | VernacStm Finish -> str"Stm Finish" + | VernacStm (Observe id) -> + str"Stm Observe " ++ str(Stateid.string_of_state_id id) + | VernacStm (Command v) -> str"Stm Command " ++ pr_vernac v + (* Proof management *) | VernacAbortAll -> str "Abort All" | VernacRestart -> str"Restart" @@ -614,7 +620,7 @@ let rec pr_vernac = function let n = List.length (List.flatten (List.map fst (List.map snd l))) in hov 2 (pr_assumption_token (n > 1) stre ++ spc() ++ - pr_ne_params_list pr_lconstr_expr l) + pr_ne_params_list pr_lconstr_expr l) | VernacInductive (f,i,l) -> let pr_constructor (coe,(id,c)) = @@ -975,3 +981,8 @@ in pr_vernac let pr_vernac_body v = make_pr_vernac pr_constr_expr pr_lconstr_expr v let pr_vernac v = make_pr_vernac pr_constr_expr pr_lconstr_expr v ++ sep_end v + +let pr_vernac x = + try pr_vernac x + with e -> Errors.print e + diff --git a/printing/prettyp.ml b/printing/prettyp.ml index 9a084adc4..0f6a5e893 100644 --- a/printing/prettyp.ml +++ b/printing/prettyp.ml @@ -420,11 +420,11 @@ let print_constant with_values sep sp = str"*** [ " ++ print_basename sp ++ str " : " ++ cut () ++ pr_ltype typ ++ str" ]" ++ - Printer.pr_univ_cstr cb.const_constraints + Printer.pr_univ_cstr (Future.force cb.const_constraints) | _ -> print_basename sp ++ str sep ++ cut () ++ (if with_values then print_typed_body (val_0,typ) else pr_ltype typ)++ - Printer.pr_univ_cstr cb.const_constraints) + Printer.pr_univ_cstr (Future.force cb.const_constraints)) let gallina_print_constant_with_infos sp = print_constant true " = " sp ++ @@ -566,7 +566,8 @@ let print_full_pure_context () = | OpaqueDef lc -> str "Theorem " ++ print_basename con ++ cut () ++ str " : " ++ pr_ltype typ ++ str "." ++ fnl () ++ - str "Proof " ++ pr_lconstr (Lazyconstr.force_opaque lc) + str "Proof " ++ pr_lconstr + (Lazyconstr.force_opaque (Future.force lc)) | Def c -> str "Definition " ++ print_basename con ++ cut () ++ str " : " ++ pr_ltype typ ++ cut () ++ str " := " ++ diff --git a/printing/printer.ml b/printing/printer.ml index 696f25695..cc9356cda 100644 --- a/printing/printer.ml +++ b/printing/printer.ml @@ -387,7 +387,7 @@ let default_pr_subgoal n sigma = | [] -> error "No such goal." | g::rest -> if Int.equal p 1 then - let pg = default_pr_goal { sigma=sigma ; it=g } in + let pg = default_pr_goal { sigma=sigma ; it=g; eff=Declareops.no_seff } in v 0 (str "subgoal " ++ int n ++ pr_goal_tag g ++ str " is:" ++ cut () ++ pg) else @@ -438,7 +438,7 @@ let default_pr_subgoals ?(pr_first=true) close_cmd sigma seeds stack goals = in let print_multiple_goals g l = if pr_first then - default_pr_goal { it = g ; sigma = sigma } ++ + default_pr_goal { it = g ; sigma = sigma; eff=Declareops.no_seff } ++ pr_rec 2 l else pr_rec 1 (g::l) @@ -463,13 +463,13 @@ let default_pr_subgoals ?(pr_first=true) close_cmd sigma seeds stack goals = str "You can use Grab Existential Variables.") end | [g],[] when not !Flags.print_emacs -> - let pg = default_pr_goal { it = g ; sigma = sigma } in + let pg = default_pr_goal { it = g ; sigma = sigma; eff=Declareops.no_seff } in v 0 ( str "1 subgoal" ++ pr_goal_tag g ++ cut () ++ pg ++ emacs_print_dependent_evars sigma seeds ) | [g],a::l when not !Flags.print_emacs -> - let pg = default_pr_goal { it = g ; sigma = sigma } in + let pg = default_pr_goal { it = g ; sigma = sigma; eff=Declareops.no_seff } in v 0 ( str "1 focused subgoal (" ++ print_unfocused a l ++ str")" ++ pr_goal_tag g ++ cut () ++ pg ++ emacs_print_dependent_evars sigma seeds @@ -555,7 +555,7 @@ let pr_goal_by_id id = ++ pr_goal gs) in try - Proof.in_proof p (fun sigma -> pr {it=g;sigma=sigma}) + Proof.in_proof p (fun sigma -> pr {it=g;sigma=sigma;eff=Declareops.no_seff}) with Not_found -> error "Invalid goal identifier." (* Elementary tactics *) diff --git a/proofs/clenvtac.ml b/proofs/clenvtac.ml index 389075c9a..7b1ccd542 100644 --- a/proofs/clenvtac.ml +++ b/proofs/clenvtac.ml @@ -113,7 +113,7 @@ let unifyTerms ?(flags=fail_quick_unif_flags) m n gls = let env = pf_env gls in let evd = create_goal_evar_defs (project gls) in let evd' = w_unify env evd CONV ~flags m n in - tclIDTAC {it = gls.it; sigma = evd'} + tclIDTAC {it = gls.it; sigma = evd'; eff = gls.eff} let unify ?(flags=fail_quick_unif_flags) m gls = let n = pf_concl gls in unifyTerms ~flags m n gls diff --git a/proofs/goal.ml b/proofs/goal.ml index 67cd41412..c6e2c2ed9 100644 --- a/proofs/goal.ml +++ b/proofs/goal.ml @@ -87,7 +87,8 @@ type subgoals = { subgoals: goal list } optimisation, since it will generaly not change during the evaluation. *) type 'a sensitive = - Environ.env -> Evd.evar_map ref -> goal -> Evd.evar_info -> 'a + Environ.env -> Evd.evar_map ref -> goal -> Evd.evar_info -> + 'a * Declareops.side_effects (* evaluates a goal sensitive value in a given goal (knowing the current evar_map). *) (* the evar_info corresponding to the goal is computed at once @@ -96,22 +97,25 @@ let eval t env defs gl = let info = content defs gl in let env = Environ.reset_with_named_context (Evd.evar_hyps info) env in let rdefs = ref defs in - let r = t env rdefs gl info in - ( r , !rdefs ) + let r, eff = t env rdefs gl info in + ( r , !rdefs, eff ) (* monadic bind on sensitive expressions *) let bind e f env rdefs goal info = - f (e env rdefs goal info) env rdefs goal info + let a, eff1 = e env rdefs goal info in + let b, eff2 = f a env rdefs goal info in + b, Declareops.union_side_effects eff1 eff2 (* monadic return on sensitive expressions *) -let return v _ _ _ _ = v +let return v e _ _ _ _ = v, e (* interpretation of "open" constr *) (* spiwack: it is a wrapper around [Constrintern.interp_open_constr]. In an ideal world, this could/should be the other way round. As of now, though, it seems at least quite useful to build tactics. *) let interp_constr cexpr env rdefs _ _ = - Constrintern.interp_constr_evars rdefs env cexpr + let c = Constrintern.interp_constr_evars rdefs env cexpr in + c, Declareops.no_seff (* Type of constr with holes used by refine. *) (* The list of evars doesn't necessarily contain all the evars in the constr, @@ -131,19 +135,18 @@ module Refinable = struct let make t env rdefs gl info = let r = ref [] in - let me = t r env rdefs gl info in - { me = me; - my_evars = !r } + let me, eff = t r env rdefs gl info in + { me = me; my_evars = !r }, eff let make_with t env rdefs gl info = let r = ref [] in - let (me,side) = t r env rdefs gl info in - { me = me ; my_evars = !r } , side + let (me,side), eff = t r env rdefs gl info in + ({ me = me ; my_evars = !r }, side), eff let mkEvar handle env typ _ rdefs _ _ = let ev = Evarutil.e_new_evar rdefs env typ in let (e,_) = destEvar ev in handle := e::!handle; - ev + ev, Declareops.no_seff (* [with_type c typ] constrains term [c] to have type [typ]. *) let with_type t typ env rdefs _ _ = @@ -157,14 +160,14 @@ module Refinable = struct Coercion.inh_conv_coerce_to (Loc.ghost) env !rdefs j typ in rdefs := new_defs; - j'.Environ.uj_val + j'.Environ.uj_val, Declareops.no_seff (* spiwack: it is not very fine grain since it solves all typeclasses holes, not only those containing the current goal, or a given term. But it seems to fit our needs so far. *) let resolve_typeclasses ?filter ?split ?(fail=false) () env rdefs _ _ = rdefs:=Typeclasses.resolve_typeclasses ?filter ?split ~fail env !rdefs; - () + (), Declareops.no_seff @@ -225,13 +228,13 @@ module Refinable = struct Pretyping.understand_tcc_evars ~flags rdefs env ~expected_type:tycon rawc in ignore(update_handle handle init_defs !rdefs); - open_constr + open_constr, Declareops.no_seff let constr_of_open_constr handle check_type (evars, c) env rdefs gl info = let delta = update_handle handle !rdefs evars in rdefs := Evd.fold (fun ev evi evd -> Evd.add evd ev evi) delta !rdefs; if check_type then with_type c (Evd.evar_concl (content !rdefs gl)) env rdefs gl info - else c + else c, Declareops.no_seff end @@ -252,7 +255,7 @@ let refine step env rdefs gl info = let subgoals = Option.List.flatten (List.map (advance !rdefs) subgoals) in - { subgoals = subgoals } + { subgoals = subgoals }, Declareops.no_seff (*** Cleaning goals ***) @@ -266,7 +269,7 @@ let clear ids env rdefs gl info = let (cleared_evar,_) = destEvar cleared_concl in let cleared_goal = descendent gl cleared_evar in rdefs := Evd.define gl.content cleared_concl !rdefs; - { subgoals = [cleared_goal] } + { subgoals = [cleared_goal] }, Declareops.no_seff let wrap_apply_to_hyp_and_dependent_on sign id f g = try Environ.apply_to_hyp_and_dependent_on sign id f g @@ -322,27 +325,27 @@ let clear_body idents env rdefs gl info = let (new_evar,_) = destEvar new_constr in let new_goal = descendent gl new_evar in rdefs := Evd.define gl.content new_constr defs'; - { subgoals = [new_goal] } + { subgoals = [new_goal] }, Declareops.no_seff (*** Sensitive primitives ***) (* [concl] is the conclusion of the current goal *) let concl _ _ _ info = - Evd.evar_concl info + Evd.evar_concl info, Declareops.no_seff (* [hyps] is the [named_context_val] representing the hypotheses of the current goal *) let hyps _ _ _ info = - Evd.evar_hyps info + Evd.evar_hyps info, Declareops.no_seff (* [env] is the current [Environ.env] containing both the environment in which the proof is ran, and the goal hypotheses *) -let env env _ _ _ = env +let env env _ _ _ = env, Declareops.no_seff (* [defs] is the [Evd.evar_map] at the current evaluation point *) let defs _ rdefs _ _ = - !rdefs + !rdefs, Declareops.no_seff (* Cf mli for more detailed comment. [null], [plus], [here] and [here_list] use internal exception [UndefinedHere] @@ -360,7 +363,7 @@ let equal { content = e1 } { content = e2 } = Int.equal e1 e2 let here goal value _ _ goal' _ = if equal goal goal' then - value + value, Declareops.no_seff else raise UndefinedHere @@ -372,7 +375,7 @@ let rec list_mem_with eq x = function let here_list goals value _ _ goal' _ = if list_mem_with equal goal' goals then - value + value, Declareops.no_seff else raise UndefinedHere @@ -392,21 +395,23 @@ let convert_hyp check (id,b,bt as d) env rdefs gl info = d) in (* Modified named context. *) - let new_hyps = - Environ.apply_to_hyp (hyps env rdefs gl info) id replace_function + let new_hyps, effs1 = + let hyps, effs = hyps env rdefs gl info in + Environ.apply_to_hyp hyps id replace_function, effs in let new_env = Environ.reset_with_named_context new_hyps env in - let new_constr = - Evarutil.e_new_evar rdefs new_env (concl env rdefs gl info) + let new_constr, effs2 = + let concl, effs = concl env rdefs gl info in + Evarutil.e_new_evar rdefs new_env concl, effs in let (new_evar,_) = destEvar new_constr in let new_goal = descendent gl new_evar in rdefs := Evd.define gl.content new_constr !rdefs; - { subgoals = [new_goal] } + { subgoals = [new_goal] }, Declareops.union_side_effects effs1 effs2 let convert_concl check cl' env rdefs gl info = let sigma = !rdefs in - let cl = concl env rdefs gl info in + let cl, effs = concl env rdefs gl info in check_typability env sigma cl'; if (not check) || Reductionops.is_conv_leq env sigma cl' cl then let new_constr = @@ -415,7 +420,7 @@ let convert_concl check cl' env rdefs gl info = let (new_evar,_) = destEvar new_constr in let new_goal = descendent gl new_evar in rdefs := Evd.define gl.content new_constr !rdefs; - { subgoals = [new_goal] } + { subgoals = [new_goal] }, effs else Errors.error "convert-concl rule passed non-converting term" @@ -428,19 +433,20 @@ let rename_hyp_sign id1 id2 sign = (fun (_,b,t) _ -> (id2,b,t)) (fun d _ -> map_named_declaration (replace_vars [id1,mkVar id2]) d) let rename_hyp id1 id2 env rdefs gl info = - let hyps = hyps env rdefs gl info in + let hyps, effs1 = hyps env rdefs gl info in if not (Names.Id.equal id1 id2) && List.mem id2 (Termops.ids_of_named_context (Environ.named_context_of_val hyps)) then Errors.error ((Names.Id.to_string id2)^" is already used."); let new_hyps = rename_hyp_sign id1 id2 hyps in let new_env = Environ.reset_with_named_context new_hyps env in - let new_concl = Vars.replace_vars [id1,mkVar id2] (concl env rdefs gl info) in + let old_concl, effs2 = concl env rdefs gl info in + let new_concl = Vars.replace_vars [id1,mkVar id2] old_concl in let new_subproof = Evarutil.e_new_evar rdefs new_env new_concl in let new_subproof = Vars.replace_vars [id2,mkVar id1] new_subproof in let (new_evar,_) = destEvar new_subproof in let new_goal = descendent gl new_evar in rdefs := Evd.define gl.content new_subproof !rdefs; - { subgoals = [new_goal] } + { subgoals = [new_goal] }, Declareops.union_side_effects effs1 effs2 (*** Additional functions ***) @@ -452,10 +458,10 @@ let rename_hyp id1 id2 env rdefs gl info = *) let rec list_map f l s = match l with - | [] -> ([],s) - | a::l -> let (a,s) = f s a in - let (l,s) = list_map f l s in - (a::l,s) + | [] -> ([],s,[]) + | a::l -> let (a,s,e) = f s a in + let (l,s,es) = list_map f l s in + (a::l,s,e::es) (* Layer to implement v8.2 tactic engine ontop of the new architecture. @@ -521,7 +527,7 @@ module V82 = struct let evi = Evd.make_evar Environ.empty_named_context_val Term.mkProp in let evk = Evarutil.new_untyped_evar () in let sigma = Evd.add Evd.empty evk evi in - { Evd.it = build evk ; Evd.sigma = sigma } + { Evd.it = build evk ; Evd.sigma = sigma; eff = Declareops.no_seff } (* Makes a goal out of an evar *) let build = build @@ -564,7 +570,7 @@ module V82 = struct let new_evi = Typeclasses.mark_unresolvable new_evi in let evk = Evarutil.new_untyped_evar () in let new_sigma = Evd.add Evd.empty evk new_evi in - { Evd.it = build evk ; sigma = new_sigma } + { Evd.it = build evk ; sigma = new_sigma; eff = Declareops.no_seff } (* Used by the compatibility layer and typeclasses *) let nf_evar sigma gl = diff --git a/proofs/goal.mli b/proofs/goal.mli index bb1c0639e..345f46abf 100644 --- a/proofs/goal.mli +++ b/proofs/goal.mli @@ -41,13 +41,15 @@ type subgoals = private { subgoals: goal list } type +'a sensitive (* evaluates a goal sensitive value in a given goal (knowing the current evar_map). *) -val eval : 'a sensitive -> Environ.env -> Evd.evar_map -> goal -> 'a * Evd.evar_map +val eval : + 'a sensitive -> Environ.env -> Evd.evar_map -> goal -> + 'a * Evd.evar_map * Declareops.side_effects (* monadic bind on sensitive expressions *) val bind : 'a sensitive -> ('a -> 'b sensitive) -> 'b sensitive (* monadic return on sensitive expressions *) -val return : 'a -> 'a sensitive +val return : 'a -> Declareops.side_effects -> 'a sensitive (* interpretation of "open" constr *) @@ -168,10 +170,10 @@ val here_list : goal list -> 'a -> 'a sensitive (* emulates List.map for functions of type [Evd.evar_map -> 'a -> 'b * Evd.evar_map] on lists of type 'a, by propagating new evar_map to next definition *) -val list_map : (Evd.evar_map -> 'a -> 'b * Evd.evar_map) -> - 'a list -> - Evd.evar_map -> - 'b list *Evd.evar_map +val list_map : + (Evd.evar_map -> 'a -> 'b * Evd.evar_map * Declareops.side_effects) -> + 'a list -> Evd.evar_map -> + 'b list * Evd.evar_map * Declareops.side_effects list (* Layer to implement v8.2 tactic engine ontop of the new architecture. Types are different from what they used to be due to a change of the diff --git a/proofs/logic.ml b/proofs/logic.ml index 7dc3fb49a..68031e2b5 100644 --- a/proofs/logic.ml +++ b/proofs/logic.ml @@ -571,13 +571,12 @@ let prim_refiner r sigma goal = mk_sign (push_named_context_val (f,None,ar) sign) oth | [] -> Goal.list_map (fun sigma (_,_,c) -> - let (gl,ev,sig')= - Goal.V82.mk_goal sigma sign c - (Goal.V82.extra sigma goal) - in ((gl,ev),sig')) - all sigma + let gl,ev,sig' = + Goal.V82.mk_goal sigma sign c (Goal.V82.extra sigma goal) in + (gl,ev),sig',Declareops.no_seff) + all sigma in - let (gls_evs,sigma) = mk_sign sign all in + let (gls_evs,sigma,_) = mk_sign sign all in let (gls,evs) = List.split gls_evs in let ids = List.map pi1 all in let evs = List.map (Vars.subst_vars (List.rev ids)) evs in @@ -612,15 +611,14 @@ let prim_refiner r sigma goal = with | Not_found -> mk_sign (push_named_context_val (f,None,ar) sign) oth) - | [] -> Goal.list_map (fun sigma(_,c) -> - let (gl,ev,sigma) = - Goal.V82.mk_goal sigma sign c - (Goal.V82.extra sigma goal) - in - ((gl,ev),sigma)) - all sigma + | [] -> + Goal.list_map (fun sigma(_,c) -> + let gl,ev,sigma = + Goal.V82.mk_goal sigma sign c (Goal.V82.extra sigma goal) in + (gl,ev),sigma,Declareops.no_seff) + all sigma in - let (gls_evs,sigma) = mk_sign sign all in + let (gls_evs,sigma,_) = mk_sign sign all in let (gls,evs) = List.split gls_evs in let (ids,types) = List.split all in let evs = List.map (Vars.subst_vars (List.rev ids)) evs in diff --git a/proofs/pfedit.ml b/proofs/pfedit.ml index 6132cc9c6..ed1de75bc 100644 --- a/proofs/pfedit.ml +++ b/proofs/pfedit.ml @@ -26,49 +26,24 @@ let delete_proof = Proof_global.discard let delete_current_proof = Proof_global.discard_current let delete_all_proofs = Proof_global.discard_all -let undo n = - let p = Proof_global.give_me_the_proof () in - let d = Proof.V82.depth p in - if n >= d then raise Proof.EmptyUndoStack; - for _i = 1 to n do - Proof.undo p - done - -let current_proof_depth () = - try - let p = Proof_global.give_me_the_proof () in - Proof.V82.depth p - with Proof_global.NoCurrentProof -> -1 - -(* [undo_todepth n] resets the proof to its nth step (does [undo (d-n)] where d - is the depth of the focus stack). *) -let undo_todepth n = - try - undo ((current_proof_depth ()) - n ) - with Proof_global.NoCurrentProof when Int.equal n 0 -> () +let set_undo _ = () +let get_undo _ = None -let start_proof id str hyps c ?init_tac ?compute_guard hook = +let start_proof (id : Id.t) str hyps c ?init_tac ?compute_guard hook = let goals = [ (Global.env_of_context hyps , c) ] in Proof_global.start_proof id str goals ?compute_guard hook; - let tac = match init_tac with - | Some tac -> Proofview.V82.tactic tac - | None -> Proofview.tclUNIT () - in - try Proof_global.run_tactic tac - with reraise -> - let reraise = Errors.push reraise in - Proof_global.discard_current (); - raise reraise - -let restart_proof () = undo_todepth 1 - -let cook_proof hook = - let prf = Proof_global.give_me_the_proof () in - hook prf; - match Proof_global.close_proof () with + let env = Global.env () in + Proof_global.with_current_proof (fun _ p -> + match init_tac with + | None -> p + | Some tac -> Proof.run_tactic env (Proofview.V82.tactic tac) p) + +let cook_this_proof hook p = + match p with | (i,([e],cg,str,h)) -> (i,(e,cg,str,h)) | _ -> Errors.anomaly ~label:"Pfedit.cook_proof" (Pp.str "more than one proof term.") +let cook_proof hook = cook_this_proof hook (Proof_global.close_proof ()) let get_pftreestate () = Proof_global.give_me_the_proof () @@ -88,15 +63,15 @@ let _ = Errors.register_handler begin function end let get_nth_V82_goal i = let p = Proof_global.give_me_the_proof () in - let { it=goals ; sigma = sigma } = Proof.V82.subgoals p in + let { it=goals ; sigma = sigma; eff = eff } = Proof.V82.subgoals p in try - { it=(List.nth goals (i-1)) ; sigma=sigma } + { it=(List.nth goals (i-1)) ; sigma=sigma; eff = eff } with Failure _ -> raise NoSuchGoal let get_goal_context_gen i = try - let { it=goal ; sigma=sigma } = get_nth_V82_goal i in - (sigma, Refiner.pf_env { it=goal ; sigma=sigma }) +let { it=goal ; sigma=sigma; eff=eff } = get_nth_V82_goal i in +(sigma, Refiner.pf_env { it=goal ; sigma=sigma; eff=eff }) with Proof_global.NoCurrentProof -> Errors.error "No focused proof." let get_goal_context i = @@ -115,26 +90,23 @@ let current_proof_statement () = | (id,([concl],strength,hook)) -> id,strength,concl,hook | _ -> Errors.anomaly ~label:"Pfedit.current_proof_statement" (Pp.str "more than one statement") -let solve_nth ?(with_end_tac=false) gi tac = +let solve_nth ?with_end_tac gi tac pr = try let tac = Proofview.V82.tactic tac in - let tac = if with_end_tac then - Proof_global.with_end_tac tac - else - tac - in - Proof_global.run_tactic (Proofview.tclFOCUS gi gi tac) - with + let tac = match with_end_tac with + | None -> tac + | Some etac -> Proofview.tclTHEN tac etac in + Proof.run_tactic (Global.env ()) (Proofview.tclFOCUS gi gi tac) pr + with | Proof_global.NoCurrentProof -> Errors.error "No focused proof" | Proofview.IndexOutOfRange -> let msg = str "No such goal: " ++ int gi ++ str "." in Errors.errorlabstrm "" msg -let by = solve_nth 1 +let by tac = Proof_global.with_current_proof (fun _ -> solve_nth 1 tac) let instantiate_nth_evar_com n com = - let pf = Proof_global.give_me_the_proof () in - Proof.V82.instantiate_evar n com pf + Proof_global.with_current_proof (fun _ -> Proof.V82.instantiate_evar n com) (**********************************************************************) @@ -144,8 +116,8 @@ open Decl_kinds let next = let n = ref 0 in fun () -> incr n; !n -let build_constant_by_tactic id sign typ tac = - start_proof id (Global,Proof Theorem) sign typ (fun _ _ -> ()); +let build_constant_by_tactic id sign ?(goal_kind = Global,Proof Theorem) typ tac = + start_proof id goal_kind sign typ (fun _ _ -> ()); try by tac; let _,(const,_,_,_) = cook_proof (fun _ -> ()) in @@ -155,10 +127,19 @@ let build_constant_by_tactic id sign typ tac = delete_current_proof (); raise reraise +let constr_of_def = function + | Declarations.Undef _ -> assert false + | Declarations.Def cs -> Lazyconstr.force cs + | Declarations.OpaqueDef lc -> Lazyconstr.force_opaque (Future.force lc) + let build_by_tactic env typ tac = let id = Id.of_string ("temporary_proof"^string_of_int (next())) in let sign = val_of_named_context (named_context env) in - (build_constant_by_tactic id sign typ tac).const_entry_body + let ce = build_constant_by_tactic id sign typ tac in + let ce = Term_typing.handle_side_effects env ce in + let cb, se = Future.force ce.const_entry_body in + assert(se = Declareops.no_seff); + cb (**********************************************************************) (* Support for resolution of evars in tactic interpretation, including diff --git a/proofs/pfedit.mli b/proofs/pfedit.mli index 976ac276e..ed9b55ae5 100644 --- a/proofs/pfedit.mli +++ b/proofs/pfedit.mli @@ -49,21 +49,6 @@ val delete_current_proof : unit -> unit val delete_all_proofs : unit -> unit -(** {6 ... } *) -(** [undo n] undoes the effect of the last [n] tactics applied to the - current proof; it fails if no proof is focused or if the ``undo'' - stack is exhausted *) - -val undo : int -> unit - -(** [undo_todepth n] resets the proof to its nth step (does [undo (d-n)] where d - is the depth of the undo stack). *) -val undo_todepth : int -> unit - -(** Returns the depth of the current focused proof stack, this is used - to put informations in coq prompt (in emacs mode). *) -val current_proof_depth: unit -> int - (** {6 ... } *) (** [start_proof s str env t hook tac] starts a proof of name [s] and conclusion [t]; [hook] is optionally a function to be applied at @@ -79,17 +64,22 @@ val start_proof : ?init_tac:tactic -> ?compute_guard:lemma_possible_guards -> unit declaration_hook -> unit -(** [restart_proof ()] restarts the current focused proof from the beginning - or fails if no proof is focused *) - -val restart_proof : unit -> unit - (** {6 ... } *) (** [cook_proof opacity] turns the current proof (assumed completed) into a constant with its name, kind and possible hook (see [start_proof]); it fails if there is no current proof of if it is not completed; it also tells if the guardness condition has to be inferred. *) +val cook_this_proof : (Proof.proof -> unit) -> + Id.t * + (Entries.definition_entry list * + lemma_possible_guards * + Decl_kinds.goal_kind * + unit Tacexpr.declaration_hook) -> + Id.t * + (Entries.definition_entry * lemma_possible_guards * goal_kind * + unit declaration_hook) + val cook_proof : (Proof.proof -> unit) -> Id.t * (Entries.definition_entry * lemma_possible_guards * goal_kind * @@ -145,7 +135,8 @@ val get_used_variables : unit -> Context.section_context option current focused proof or raises a UserError if no proof is focused or if there is no [n]th subgoal *) -val solve_nth : ?with_end_tac:bool -> int -> tactic -> unit +val solve_nth : ?with_end_tac:unit Proofview.tactic -> int -> tactic -> + Proof.proof -> Proof.proof (** [by tac] applies tactic [tac] to the 1st subgoal of the current focused proof or raises a UserError if there is no focused proof or @@ -162,8 +153,9 @@ val instantiate_nth_evar_com : int -> Constrexpr.constr_expr -> unit (** [build_by_tactic typ tac] returns a term of type [typ] by calling [tac] *) -val build_constant_by_tactic : Id.t -> named_context_val -> types -> tactic -> - Entries.definition_entry +val build_constant_by_tactic : + Id.t -> named_context_val -> ?goal_kind:goal_kind -> + types -> tactic -> Entries.definition_entry val build_by_tactic : env -> types -> tactic -> constr (** Declare the default tactic to fill implicit arguments *) diff --git a/proofs/proof.ml b/proofs/proof.ml index 6cc43de72..38886e9e1 100644 --- a/proofs/proof.ml +++ b/proofs/proof.ml @@ -38,7 +38,7 @@ type unfocusable = | Cannot of exn | Loose | Strict -type _focus_condition = +type _focus_condition = (_focus_kind -> Proofview.proofview -> unfocusable) * (_focus_kind -> bool) type 'a focus_condition = _focus_condition @@ -83,11 +83,11 @@ module Cond = struct else c let (&&&) c1 c2 k p= match c1 k p , c2 k p with - | Cannot e , _ + | Cannot e , _ | _ , Cannot e -> Cannot e | Strict, Strict -> Strict | _ , _ -> Loose - let kind e k0 k p = bool e (Int.equal k0 k) k p + let kind e k0 k p = bool e (Int.equal k0 k) k p let pdone e k p = bool e (Proofview.finished p) k p end @@ -116,7 +116,7 @@ let done_cond ?loose_end = done_cond_gen CannotUnfocusThisWay ?loose_end (* Subpart of the type of proofs. It contains the parts of the proof which are under control of the undo mechanism *) -type proof_state = { +type proof = { (* Current focused proofview *) proofview: Proofview.proofview; (* History of the focusings, provides information on how @@ -125,29 +125,9 @@ type proof_state = { focus_stack: (_focus_condition*focus_info*Proofview.focus_context) list; } -type proof_info = { - mutable endline_tactic : unit Proofview.tactic ; - mutable section_vars : Context.section_context option; - initial_conclusions : Term.types list -} - -type undo_action = - | State of proof_state - | Effect of (unit -> unit) - -type proof = { (* current proof_state *) - mutable state : proof_state; - (* The undo stack *) - mutable undo_stack : undo_action list; - (* secondary undo stacks used for transactions *) - mutable transactions : undo_action list list; - info : proof_info - } - - (*** General proof functions ***) -let proof { state = p } = +let proof p = let (goals,sigma) = Proofview.proofview p.proofview in (* spiwack: beware, the bottom of the stack is used by [Proof] internally, and should not be exposed. *) @@ -169,25 +149,22 @@ let rec unroll_focus pv = function doesn't hide any goal. Unfocusing is handled in {!return}. *) let is_done p = - Proofview.finished p.state.proofview && - Proofview.finished (unroll_focus p.state.proofview p.state.focus_stack) + Proofview.finished p.proofview && + Proofview.finished (unroll_focus p.proofview p.focus_stack) (* spiwack: for compatibility with <= 8.2 proof engine *) let has_unresolved_evar p = - Proofview.V82.has_unresolved_evar p.state.proofview + Proofview.V82.has_unresolved_evar p.proofview (* Returns the list of partial proofs to initial goals *) -let partial_proof p = - List.map fst (Proofview.return p.state.proofview) - - +let partial_proof p = Proofview.partial_proof p.proofview (*** The following functions implement the basic internal mechanisms of proofs, they are not meant to be exported in the .mli ***) (* An auxiliary function to push a {!focus_context} on the focus stack. *) let push_focus cond inf context pr = - pr.state <- { pr.state with focus_stack = (cond,inf,context)::pr.state.focus_stack } + { pr with focus_stack = (cond,inf,context)::pr.focus_stack } exception FullyUnfocused let _ = Errors.register_handler begin function @@ -196,162 +173,52 @@ let _ = Errors.register_handler begin function end (* An auxiliary function to read the kind of the next focusing step *) let cond_of_focus pr = - match pr.state.focus_stack with + match pr.focus_stack with | (cond,_,_)::_ -> cond | _ -> raise FullyUnfocused (* An auxiliary function to pop and read the last {!Proofview.focus_context} on the focus stack. *) let pop_focus pr = - match pr.state.focus_stack with - | focus::other_focuses -> - pr.state <- { pr.state with focus_stack = other_focuses }; focus - | _ -> + match pr.focus_stack with + | focus::other_focuses -> + { pr with focus_stack = other_focuses }, focus + | _ -> raise FullyUnfocused -(* Auxiliary function to push a [proof_state] onto the undo stack. *) -let push_undo save pr = - match pr.transactions with - | [] -> pr.undo_stack <- save::pr.undo_stack - | stack::trans' -> pr.transactions <- (save::stack)::trans' - -(* Auxiliary function to pop and read a [save_state] from the undo stack. *) -exception EmptyUndoStack -let _ = Errors.register_handler begin function - | EmptyUndoStack -> Errors.error "Cannot undo: no more undo information" - | _ -> raise Errors.Unhandled -end -let pop_undo pr = - match pr.transactions , pr.undo_stack with - | [] , state::stack -> pr.undo_stack <- stack; state - | (state::stack)::trans', _ -> pr.transactions <- stack::trans'; state - | _ -> raise EmptyUndoStack - - (* This function focuses the proof [pr] between indices [i] and [j] *) let _focus cond inf i j pr = - let (focused,context) = Proofview.focus i j pr.state.proofview in - push_focus cond inf context pr; - pr.state <- { pr.state with proofview = focused } + let focused, context = Proofview.focus i j pr.proofview in + let pr = push_focus cond inf context pr in + { pr with proofview = focused } (* This function unfocuses the proof [pr], it raises [FullyUnfocused], if the proof is already fully unfocused. This function does not care about the condition of the current focus. *) let _unfocus pr = - let (_,_,fc) = pop_focus pr in - pr.state <- { pr.state with proofview = Proofview.unfocus fc pr.state.proofview } - - -let set_used_variables l p = - p.info.section_vars <- Some l - -let get_used_variables p = p.info.section_vars - -(*** Endline tactic ***) - -(* spiwack this is an information about the UI, it might be a good idea to move - it to the Proof_global module *) - -(* Sets the tactic to be used when a tactic line is closed with [...] *) -let set_endline_tactic tac p = - p.info.endline_tactic <- tac - -let with_end_tac pr tac = - Proofview.tclTHEN tac pr.info.endline_tactic - -(*** The following functions define the safety mechanism of the - proof system, they may be unsafe if not used carefully. There is - currently no reason to export them in the .mli ***) - -(* This functions saves the current state into a [proof_state]. *) -let save_state { state = ps } = State ps - -(* This function stores the current proof state in the undo stack. *) -let save pr = - push_undo (save_state pr) pr - -(* This function restores a state, presumably from the top of the undo stack. *) -let restore_state save pr = - match save with - | State save -> pr.state <- save - | Effect undo -> undo () - -(* Interpretes the Undo command. *) -let undo pr = - (* On a single line, since the effects commute *) - restore_state (pop_undo pr) pr - -(* Adds an undo effect to the undo stack. Use it with care, errors here might result - in inconsistent states. *) -let add_undo effect pr = - push_undo (Effect effect) pr - - - -(*** Transactions ***) - -let init_transaction pr = - pr.transactions <- []::pr.transactions - -let commit_stack pr stack = - let push s = push_undo s pr in - List.iter push (List.rev stack) - -(* Invariant: [commit] should be called only when a transaction - is open. It closes the current transaction. *) -let commit pr = - match pr.transactions with - | stack::trans' -> - pr.transactions <- trans'; - commit_stack pr stack - | [] -> assert false - -(* Invariant: [rollback] should be called only when a transaction - is open. It closes the current transaction. *) -let rec rollback pr = - try - undo pr; - rollback pr - with EmptyUndoStack -> - match pr.transactions with - | []::trans' -> pr.transactions <- trans' - | _ -> assert false - - -let transaction pr t = - init_transaction pr; - try t (); commit pr - with reraise -> - let reraise = Errors.push reraise in - rollback pr; raise reraise - + let pr, (_,_,fc) = pop_focus pr in + { pr with proofview = Proofview.unfocus fc pr.proofview } (* Focus command (focuses on the [i]th subgoal) *) (* spiwack: there could also, easily be a focus-on-a-range tactic, is there a need for it? *) let focus cond inf i pr = - save pr; _focus cond (Obj.repr inf) i i pr let rec unfocus kind pr () = - let starting_point = save_state pr in let cond = cond_of_focus pr in - match fst cond kind pr.state.proofview with + match fst cond kind pr.proofview with | Cannot e -> raise e - | Strict -> - (_unfocus pr; - push_undo starting_point pr) + | Strict -> + let pr = _unfocus pr in + pr | Loose -> begin try - _unfocus pr; - push_undo starting_point pr; + let pr = _unfocus pr in unfocus kind pr () with FullyUnfocused -> raise CannotUnfocusThisWay end -let unfocus kind pr = - transaction pr (unfocus kind pr) - exception NoSuchFocus (* no handler: should not be allowed to reach toplevel. *) let rec get_in_focus_stack kind stack = @@ -361,14 +228,20 @@ let rec get_in_focus_stack kind stack = else get_in_focus_stack kind stack | [] -> raise NoSuchFocus let get_at_focus kind pr = - Obj.magic (get_in_focus_stack kind pr.state.focus_stack) + Obj.magic (get_in_focus_stack kind pr.focus_stack) let is_last_focus kind pr = - let ((_,check),_,_) = List.hd pr.state.focus_stack in + let ((_,check),_,_) = List.hd pr.focus_stack in check kind let no_focused_goal p = - Proofview.finished p.state.proofview + Proofview.finished p.proofview + +let rec maximal_unfocus k p = + if no_focused_goal p then + try maximal_unfocus k (unfocus k p ()) + with FullyUnfocused | CannotUnfocusThisWay -> p + else p (*** Proof Creation/Termination ***) @@ -379,18 +252,10 @@ let end_of_stack = done_cond_gen FullyUnfocused end_of_stack_kind let unfocused = is_last_focus end_of_stack_kind let start goals = - let pr = - { state = { proofview = Proofview.init goals ; - focus_stack = [] }; - undo_stack = [] ; - transactions = [] ; - info = { endline_tactic = Proofview.tclUNIT (); - initial_conclusions = List.map snd goals; - section_vars = None } - } - in - _focus end_of_stack (Obj.repr ()) 1 (List.length goals) pr; - pr + let pr = { + proofview = Proofview.init goals ; + focus_stack = [] } in + _focus end_of_stack (Obj.repr ()) 1 (List.length goals) pr exception UnfinishedProof exception HasUnresolvedEvar @@ -399,73 +264,64 @@ let _ = Errors.register_handler begin function | HasUnresolvedEvar -> Errors.error "Some existential variables are uninstantiated." | _ -> raise Errors.Unhandled end -let return p = + +let return p t = if not (is_done p) then raise UnfinishedProof - else if has_unresolved_evar p then + else if has_unresolved_evar p then (* spiwack: for compatibility with <= 8.3 proof engine *) raise HasUnresolvedEvar else - unfocus end_of_stack_kind p; - Proofview.return p.state.proofview + let p = unfocus end_of_stack_kind p () in + Proofview.return p.proofview t + +let initial_goals p = Proofview.initial_goals p.proofview + +(*** Function manipulation proof extra informations ***) (*** Tactics ***) let run_tactic env tac pr = - let starting_point = save_state pr in - let sp = pr.state.proofview in - try - let tacticced_proofview = Proofview.apply env tac sp in - pr.state <- { pr.state with proofview = tacticced_proofview }; - push_undo starting_point pr - with reraise -> - restore_state starting_point pr; - raise reraise + let sp = pr.proofview in + let tacticced_proofview = Proofview.apply env tac sp in + { pr with proofview = tacticced_proofview } -(*** Commands ***) +let emit_side_effects eff pr = + {pr with proofview = Proofview.emit_side_effects eff pr.proofview} -let in_proof p k = - Proofview.in_proofview p.state.proofview k +(*** Commands ***) +let in_proof p k = Proofview.in_proofview p.proofview k (*** Compatibility layer with <=v8.2 ***) module V82 = struct let subgoals p = - Proofview.V82.goals p.state.proofview + Proofview.V82.goals p.proofview let background_subgoals p = - Proofview.V82.goals (unroll_focus p.state.proofview p.state.focus_stack) - - let get_initial_conclusions p = - p.info.initial_conclusions + Proofview.V82.goals (unroll_focus p.proofview p.focus_stack) - let depth p = List.length p.undo_stack - - let top_goal p = - let { Evd.it=gls ; sigma=sigma } = - Proofview.V82.top_goals p.state.proofview + let top_goal p = + let { Evd.it=gls ; sigma=sigma; eff=eff } = + Proofview.V82.top_goals p.proofview in - { Evd.it=List.hd gls ; sigma=sigma } + { Evd.it=List.hd gls ; sigma=sigma; eff=eff } let top_evars p = - Proofview.V82.top_evars p.state.proofview + Proofview.V82.top_evars p.proofview let grab_evars p = if not (is_done p) then raise UnfinishedProof else - save p; - p.state <- { p.state with proofview = Proofview.V82.grab p.state.proofview } - + { p with proofview = Proofview.V82.grab p.proofview } + let instantiate_evar n com pr = - let starting_point = save_state pr in - let sp = pr.state.proofview in + let sp = pr.proofview in try let new_proofview = Proofview.V82.instantiate_evar n com sp in - pr.state <- { pr.state with proofview = new_proofview }; - push_undo starting_point pr - with reraise -> - restore_state starting_point pr; - raise reraise + { pr with proofview = new_proofview } + with e -> + raise e end diff --git a/proofs/proof.mli b/proofs/proof.mli index b563452a9..70615e09e 100644 --- a/proofs/proof.mli +++ b/proofs/proof.mli @@ -47,6 +47,7 @@ val proof : proof -> Goal.goal list * (Goal.goal list * Goal.goal list) list * E (*** General proof functions ***) val start : (Environ.env * Term.types) list -> proof +val initial_goals : proof -> (Term.constr * Term.types) list (* Returns [true] if the considered proof is completed, that is if no goal remain to be considered (this does not require that all evars have been solved). *) @@ -60,20 +61,8 @@ val partial_proof : proof -> Term.constr list Raises [HasUnresolvedEvar] if some evars have been left undefined. *) exception UnfinishedProof exception HasUnresolvedEvar -val return : proof -> (Term.constr * Term.types) list - -(* Interpretes the Undo command. Raises [EmptyUndoStack] if - the undo stack is empty. *) -exception EmptyUndoStack -val undo : proof -> unit - -(* Adds an undo effect to the undo stack. Use it with care, errors here might result - in inconsistent states. - An undo effect is meant to undo an effect on a proof (a canonical example - of which is {!Proofglobal.set_proof_mode} which changes the current parser for - tactics). Make sure it will work even if the effects have been only partially - applied at the time of failure. *) -val add_undo : (unit -> unit) -> proof -> unit +val return : + proof -> Term.constr -> Term.constr * Declareops.side_effects (*** Focusing actions ***) @@ -113,7 +102,7 @@ val done_cond : ?loose_end:bool -> 'a focus_kind -> 'a focus_condition (* focus command (focuses on the [i]th subgoal) *) (* spiwack: there could also, easily be a focus-on-a-range tactic, is there a need for it? *) -val focus : 'a focus_condition -> 'a -> int -> proof -> unit +val focus : 'a focus_condition -> 'a -> int -> proof -> proof exception FullyUnfocused exception CannotUnfocusThisWay @@ -121,7 +110,7 @@ exception CannotUnfocusThisWay Raises [FullyUnfocused] if the proof is not focused. Raises [CannotUnfocusThisWay] if the proof the unfocusing condition is not met. *) -val unfocus : 'a focus_kind -> proof -> unit +val unfocus : 'a focus_kind -> proof -> unit -> proof (* [unfocused p] returns [true] when [p] is fully unfocused. *) val unfocused : proof -> bool @@ -138,34 +127,13 @@ val is_last_focus : 'a focus_kind -> proof -> bool (* returns [true] if there is no goal under focus. *) val no_focused_goal : proof -> bool -(* Sets the section variables assumed by the proof *) -val set_used_variables : Context.section_context -> proof -> unit -val get_used_variables : proof -> Context.section_context option - -(*** Endline tactic ***) - -(* Sets the tactic to be used when a tactic line is closed with [...] *) -val set_endline_tactic : unit Proofview.tactic -> proof -> unit - -val with_end_tac : proof -> unit Proofview.tactic -> unit Proofview.tactic - (*** Tactics ***) -val run_tactic : Environ.env -> unit Proofview.tactic -> proof -> unit - +val run_tactic : Environ.env -> unit Proofview.tactic -> proof -> proof -(*** Transactions ***) - -(* A transaction chains several commands into a single one. For instance, - a focusing command and a tactic. Transactions are such that if - any of the atomic action fails, the whole transaction fails. - - During a transaction, the visible undo stack is constituted only - of the actions performed done during the transaction. - - [transaction p f] can be called on an [f] using, itself, [transaction p].*) -val transaction : proof -> (unit -> unit) -> unit +val emit_side_effects : Declareops.side_effects -> proof -> proof +val maximal_unfocus : 'a focus_kind -> proof -> proof (*** Commands ***) @@ -178,19 +146,15 @@ module V82 : sig (* All the subgoals of the proof, including those which are not focused. *) val background_subgoals : proof -> Goal.goal list Evd.sigma - val get_initial_conclusions : proof -> Term.types list - - val depth : proof -> int - val top_goal : proof -> Goal.goal Evd.sigma - + (* returns the existential variable used to start the proof *) val top_evars : proof -> Evd.evar list (* Turns the unresolved evars into goals. Raises [UnfinishedProof] if there are still unsolved goals. *) - val grab_evars : proof -> unit + val grab_evars : proof -> proof (* Implements the Existential command *) - val instantiate_evar : int -> Constrexpr.constr_expr -> proof -> unit + val instantiate_evar : int -> Constrexpr.constr_expr -> proof -> proof end diff --git a/proofs/proof_global.ml b/proofs/proof_global.ml index ed985f292..0c0aac715 100644 --- a/proofs/proof_global.ml +++ b/proofs/proof_global.ml @@ -17,6 +17,7 @@ open Util open Pp open Names +open Util (*** Proof Modes ***) @@ -59,32 +60,31 @@ let _ = (*** Proof Global Environment ***) -(* local shorthand *) -type nproof = Id.t*Proof.proof - (* Extra info on proofs. *) type lemma_possible_guards = int list list -type proof_info = { - strength : Decl_kinds.goal_kind ; - compute_guard : lemma_possible_guards; - hook : unit Tacexpr.declaration_hook ; - mode : proof_mode + +type pstate = { + pid : Id.t; + endline_tactic : unit Proofview.tactic; + section_vars : Context.section_context option; + proof : Proof.proof; + strength : Decl_kinds.goal_kind; + compute_guard : lemma_possible_guards; + hook : unit Tacexpr.declaration_hook; + mode : proof_mode; } -(* Invariant: the domain of proof_info is current_proof.*) -(* The head of [!current_proof] is the actual current proof, the other ones are +(* The head of [!pstates] is the actual current proof, the other ones are to be resumed when the current proof is closed or aborted. *) -let current_proof = ref ([]:nproof list) -let proof_info = ref (Id.Map.empty : proof_info Id.Map.t) +let pstates = ref ([] : pstate list) (* Current proof_mode, for bookkeeping *) let current_proof_mode = ref !default_proof_mode (* combinators for proof modes *) let update_proof_mode () = - match !current_proof with - | (id,_)::_ -> - let { mode = m } = Id.Map.find id !proof_info in + match !pstates with + | { mode = m } :: _ -> !current_proof_mode.reset (); current_proof_mode := m; !current_proof_mode.set () @@ -103,7 +103,7 @@ let _ = Errors.register_handler begin function end let extract id l = let rec aux = function - | ((id',_) as np)::l when Id.equal id id' -> (np,l) + | ({pid = id' } as np)::l when Id.equal id id' -> (np,l) | np::l -> let (np', l) = aux l in (np' , np::l) | [] -> raise NoSuchProof in @@ -121,10 +121,6 @@ let extract_top l = match !l with | np::l' -> l := l' ; update_proof_mode (); np | [] -> raise NoCurrentProof -let find_top l = - match !l with - | np::_ -> np - | [] -> raise NoCurrentProof (* combinators for the proof_info map *) let add id info m = @@ -135,13 +131,28 @@ let remove id m = (*** Proof Global manipulation ***) let get_all_proof_names () = - List.map fst !current_proof + List.map (function { pid = id } -> id) !pstates -let give_me_the_proof () = - snd (find_top current_proof) +let cur_pstate () = + match !pstates with + | np::_ -> np + | [] -> raise NoCurrentProof -let get_current_proof_name () = - fst (find_top current_proof) +let give_me_the_proof () = (cur_pstate ()).proof +let get_current_proof_name () = (cur_pstate ()).pid + +let with_current_proof f = + match !pstates with + | [] -> raise NoCurrentProof + | p :: rest -> + let p = { p with proof = f p.endline_tactic p.proof } in + pstates := p :: rest + +(* Sets the tactic to be used when a tactic line is closed with [...] *) +let set_endline_tactic tac = + match !pstates with + | [] -> raise NoCurrentProof + | p :: rest -> pstates := { p with endline_tactic = tac } :: rest (* spiwack: it might be considered to move error messages away. Or else to remove special exceptions from Proof_global. @@ -155,7 +166,7 @@ let msg_proofs () = | l -> (str"." ++ fnl () ++ str"Proofs currently edited:" ++ spc () ++ (pr_sequence Nameops.pr_id l) ++ str".") -let there_is_a_proof () = not (List.is_empty !current_proof) +let there_is_a_proof () = not (List.is_empty !pstates) let there_are_pending_proofs () = there_is_a_proof () let check_no_pending_proof () = if not (there_are_pending_proofs ()) then @@ -167,51 +178,40 @@ let check_no_pending_proof () = end let discard_gen id = - ignore (extract id current_proof); - remove id proof_info + pstates := List.filter (fun { pid = id' } -> id <> id') !pstates let discard (loc,id) = - try - discard_gen id - with NoSuchProof -> + let n = List.length !pstates in + discard_gen id; + if List.length !pstates = n then Errors.user_err_loc (loc,"Pfedit.delete_proof",str"No such proof" ++ msg_proofs ()) let discard_current () = - let (id,_) = extract_top current_proof in - remove id proof_info + if !pstates = [] then raise NoCurrentProof else pstates := List.tl !pstates -let discard_all () = - current_proof := []; - proof_info := Id.Map.empty +let discard_all () = pstates := [] (* [set_proof_mode] sets the proof mode to be used after it's called. It is typically called by the Proof Mode command. *) -(* Core component. - No undo handling. - Applies to proof [id], and proof mode [m]. *) let set_proof_mode m id = - let info = Id.Map.find id !proof_info in - let info = { info with mode = m } in - proof_info := Id.Map.add id info !proof_info; + pstates := + List.map (function { pid = id' } as p -> + if Id.equal id' id then { p with mode = m } else p) !pstates; update_proof_mode () -(* Complete function. - Handles undo. - Applies to current proof, and proof mode name [mn]. *) + let set_proof_mode mn = - let m = find_proof_mode mn in - let id = get_current_proof_name () in - let pr = give_me_the_proof () in - Proof.add_undo begin let curr = !current_proof_mode in fun () -> - set_proof_mode curr id ; update_proof_mode () - end pr ; - set_proof_mode m id + set_proof_mode (find_proof_mode mn) (get_current_proof_name ()) + +let activate_proof_mode mode = (find_proof_mode mode).set () +let disactivate_proof_mode mode = (find_proof_mode mode).reset () exception AlreadyExists let _ = Errors.register_handler begin function | AlreadyExists -> Errors.error "Already editing something of that name." | _ -> raise Errors.Unhandled end + (* [start_proof s str env t hook tac] starts a proof of name [s] and conclusion [t]; [hook] is optionally a function to be applied at proof end (e.g. to declare the built constructions as a coercion @@ -220,84 +220,67 @@ end proof of mutually dependent theorems). It raises exception [ProofInProgress] if there is a proof being currently edited. *) -let start_proof id str goals ?(compute_guard=[]) hook = - begin - List.iter begin fun (id_ex,_) -> - if Names.Id.equal id id_ex then raise AlreadyExists - end !current_proof - end; - let p = Proof.start goals in - add id { strength=str ; - compute_guard=compute_guard ; - hook=hook ; - mode = ! default_proof_mode } proof_info ; - push (id,p) current_proof - -(* arnaud: à enlever *) -let run_tactic tac = - let p = give_me_the_proof () in - let env = Global.env () in - Proof.run_tactic env tac p -(* Sets the tactic to be used when a tactic line is closed with [...] *) -let set_endline_tactic tac = - let p = give_me_the_proof () in - Proof.set_endline_tactic tac p +let start_proof id str goals ?(compute_guard=[]) hook = + let initial_state = { + pid = id; + proof = Proof.start goals; + endline_tactic = Proofview.tclUNIT (); + section_vars = None; + strength = str; + compute_guard = compute_guard; + hook = hook; + mode = ! default_proof_mode } in + push initial_state pstates + +let get_used_variables () = (cur_pstate ()).section_vars let set_used_variables l = - let p = give_me_the_proof () in let env = Global.env () in let ids = List.fold_right Id.Set.add l Id.Set.empty in let ctx = Environ.keep_hyps env ids in - Proof.set_used_variables ctx p - -let get_used_variables () = - Proof.get_used_variables (give_me_the_proof ()) - -let with_end_tac tac = - let p = give_me_the_proof () in - Proof.with_end_tac p tac - + match !pstates with + | [] -> raise NoCurrentProof + | p :: rest -> + if p.section_vars <> None then + Errors.error "Used section variables can be declared only once"; + pstates := { p with section_vars = Some ctx} :: rest + +let get_open_goals () = + let gl, gll, _ = Proof.proof (cur_pstate ()).proof in + List.length gl + + List.fold_left (+) 0 + (List.map (fun (l1,l2) -> List.length l1 + List.length l2) gll) + +type closed_proof = + Names.Id.t * + (Entries.definition_entry list * lemma_possible_guards * + Decl_kinds.goal_kind * unit Tacexpr.declaration_hook) + +let close_proof ~now ?(fix_exn = fun x -> x) ps side_eff = + let { pid; section_vars; compute_guard; strength; hook; proof } = ps in + let entries = List.map (fun (c, t) -> { Entries. + const_entry_body = Future.chain side_eff (fun () -> + try Proof.return (cur_pstate ()).proof c with + | Proof.UnfinishedProof -> + raise (fix_exn(Errors.UserError("last tactic before Qed", + str"Attempt to save an incomplete proof"))) + | Proof.HasUnresolvedEvar -> + raise (fix_exn(Errors.UserError("last tactic before Qed", + str"Attempt to save a proof with existential " ++ + str"variables still non-instantiated")))); + const_entry_secctx = section_vars; + const_entry_type = Some t; + const_entry_inline_code = false; + const_entry_opaque = true }) (Proof.initial_goals proof) in + if now then + List.iter (fun x -> ignore(Future.join x.Entries.const_entry_body)) entries; + (pid, (entries, compute_guard, strength, hook)) + +let close_future_proof ~fix_exn proof = + close_proof ~now:false ~fix_exn (cur_pstate ()) proof let close_proof () = - (* spiwack: for now close_proof doesn't actually discard the proof, it is done - by [Command.save]. *) - try - let id = get_current_proof_name () in - let p = give_me_the_proof () in - let proofs_and_types = Proof.return p in - let section_vars = Proof.get_used_variables p in - let entries = List.map - (fun (c,t) -> { Entries.const_entry_body = c; - const_entry_secctx = section_vars; - const_entry_type = Some t; - const_entry_opaque = true; - const_entry_inline_code = false }) - proofs_and_types - in - let { compute_guard=cg ; strength=str ; hook=hook } = - Id.Map.find id !proof_info - in - (id, (entries,cg,str,hook)) - with - | Proof.UnfinishedProof -> - Errors.error "Attempt to save an incomplete proof" - | Proof.HasUnresolvedEvar -> - Errors.error "Attempt to save a proof with existential variables still non-instantiated" - - -(**********************************************************) -(* *) -(* Utility functions *) -(* *) -(**********************************************************) - -let maximal_unfocus k p = - begin try while Proof.no_focused_goal p do - Proof.unfocus k p - done - with Proof.FullyUnfocused | Proof.CannotUnfocusThisWay -> () - end - + close_proof ~now:true (cur_pstate ()) (Future.from_val()) (**********************************************************) (* *) @@ -317,7 +300,7 @@ module Bullet = struct type behavior = { name : string; - put : Proof.proof -> t -> unit + put : Proof.proof -> t -> Proof.proof } let behaviors = Hashtbl.create 4 @@ -326,7 +309,7 @@ module Bullet = struct (*** initial modes ***) let none = { name = "None"; - put = fun _ _ -> () + put = fun x _ -> x } let _ = register_behavior none @@ -356,8 +339,8 @@ module Bullet = struct let pop pr = match get_bullets pr with | b::_ -> - Proof.unfocus bullet_kind pr; - (*returns*) b + let pr = Proof.unfocus bullet_kind pr () in + pr, b | _ -> assert false let push b pr = @@ -365,10 +348,10 @@ module Bullet = struct let put p bul = if has_bullet bul p then - Proof.transaction p begin fun () -> - while not (bullet_eq bul (pop p)) do () done; - push bul p - end + let rec aux p = + let p, b = pop p in + if not (bullet_eq bul b) then aux p else p in + push bul (aux p) else push bul p @@ -403,11 +386,11 @@ end module V82 = struct let get_current_initial_conclusions () = - let p = give_me_the_proof () in - let id = get_current_proof_name () in - let { strength=str ; hook=hook } = - Id.Map.find id !proof_info - in - (id,(Proof.V82.get_initial_conclusions p, str, hook)) + let { pid; strength; hook; proof } = cur_pstate () in + pid, (List.map snd (Proof.initial_goals proof), strength, hook) end +type state = pstate list +let freeze () = !pstates +let unfreeze s = pstates := s; update_proof_mode () + diff --git a/proofs/proof_global.mli b/proofs/proof_global.mli index 7fa44cf86..a5fe33992 100644 --- a/proofs/proof_global.mli +++ b/proofs/proof_global.mli @@ -28,7 +28,6 @@ type proof_mode = { It corresponds to Coq default setting are they are set when coqtop starts. *) val register_proof_mode : proof_mode -> unit -val there_is_a_proof : unit -> bool val there_are_pending_proofs : unit -> bool val check_no_pending_proof : unit -> unit @@ -60,18 +59,27 @@ val start_proof : Names.Id.t -> unit Tacexpr.declaration_hook -> unit -val close_proof : unit -> - Names.Id.t * - (Entries.definition_entry list * - lemma_possible_guards * - Decl_kinds.goal_kind * - unit Tacexpr.declaration_hook) +type closed_proof = + Names.Id.t * + (Entries.definition_entry list * lemma_possible_guards * + Decl_kinds.goal_kind * unit Tacexpr.declaration_hook) + +val close_proof : unit -> closed_proof + +(* A future proof may be executed later on, out of the control of Stm + that knows which state was (for example) supposed to close the proof + bit did not. Hence fix_exn to enrich it *) +val close_future_proof : + fix_exn:(exn -> exn) -> unit Future.computation -> closed_proof exception NoSuchProof +val get_open_goals : unit -> int + (** Runs a tactic on the current proof. Raises [NoCurrentProof] is there is no current proof. *) -val run_tactic : unit Proofview.tactic -> unit +val with_current_proof : + (unit Proofview.tactic -> Proof.proof -> Proof.proof) -> unit (** Sets the tactic to be used when a tactic line is closed with [...] *) val set_endline_tactic : unit Proofview.tactic -> unit @@ -80,18 +88,19 @@ val set_endline_tactic : unit Proofview.tactic -> unit val set_used_variables : Names.Id.t list -> unit val get_used_variables : unit -> Context.section_context option -(** Appends the endline tactic of the current proof to a tactic. *) -val with_end_tac : unit Proofview.tactic -> unit Proofview.tactic +val discard : Names.identifier Loc.located -> unit +val discard_current : unit -> unit +val discard_all : unit -> unit (**********************************************************) -(* *) -(* Utility functions *) -(* *) +(* *) +(* Proof modes *) +(* *) (**********************************************************) -(** [maximal_unfocus k p] unfocuses [p] until [p] has at least a - focused goal or that the last focus isn't of kind [k]. *) -val maximal_unfocus : 'a Proof.focus_kind -> Proof.proof -> unit + +val activate_proof_mode : string -> unit +val disactivate_proof_mode : string -> unit (**********************************************************) (* *) @@ -107,7 +116,7 @@ module Bullet : sig with a name to identify it. *) type behavior = { name : string; - put : Proof.proof -> t -> unit + put : Proof.proof -> t -> Proof.proof } (** A registered behavior can then be accessed in Coq @@ -123,9 +132,13 @@ module Bullet : sig (** Handles focusing/defocusing with bullets: *) - val put : Proof.proof -> t -> unit + val put : Proof.proof -> t -> Proof.proof end module V82 : sig val get_current_initial_conclusions : unit -> Names.Id.t *(Term.types list * Decl_kinds.goal_kind * unit Tacexpr.declaration_hook) end + +type state +val freeze : unit -> state +val unfreeze : state -> unit diff --git a/proofs/proofview.ml b/proofs/proofview.ml index 45b7c084d..8a2c7617c 100644 --- a/proofs/proofview.ml +++ b/proofs/proofview.ml @@ -26,7 +26,8 @@ open Util type proofview = { initial : (Term.constr * Term.types) list; solution : Evd.evar_map; - comb : Goal.goal list + comb : Goal.goal list; + side_effects : Declareops.side_effects } let proofview p = @@ -38,7 +39,8 @@ let init = let rec aux = function | [] -> { initial = [] ; solution = Evd.empty ; - comb = [] + comb = []; + side_effects = Declareops.no_seff } | (env,typ)::l -> let { initial = ret ; solution = sol ; comb = comb } = aux l @@ -50,12 +52,15 @@ let init = let gl = Goal.build e in { initial = (econstr,typ)::ret; solution = new_defs ; - comb = gl::comb } + comb = gl::comb; + side_effects = Declareops.no_seff } in fun l -> let v = aux l in (* Marks all the goal unresolvable for typeclasses. *) { v with solution = Typeclasses.mark_unresolvables v.solution } +let initial_goals { initial } = initial + (* Returns whether this proofview is finished or not. That is, if it has empty subgoals in the comb. There could still be unsolved subgoaled, but they would then be out of the view, focused out. *) @@ -64,8 +69,13 @@ let finished = function | _ -> false (* Returns the current value of the proofview partial proofs. *) -let return { initial=init; solution=defs } = - List.map (fun (c,t) -> (Evarutil.nf_evar defs c , t)) init +let return { solution=defs; side_effects=eff } c = Evarutil.nf_evar defs c, eff + +let partial_proof pv = + List.map fst (List.map (return pv) (List.map fst (initial_goals pv))) + +let emit_side_effects eff x = + { x with side_effects = Declareops.union_side_effects eff x.side_effects } (* spiwack: this function should probably go in the Util section, but I'd rather have Util (or a separate module for lists) @@ -356,8 +366,8 @@ let rec catchable_exception = function [s] to each goal successively and applying [k] to each result. *) let list_of_sensitive s k env step = Goal.list_map begin fun defs g -> - let (a,defs) = Goal.eval s env defs g in - (k a) , defs + let (a,defs,effs) = Goal.eval s env defs g in + (k a) , defs, effs end step.comb step.solution (* In form of a tactic *) let list_of_sensitive s k env = @@ -366,8 +376,10 @@ let list_of_sensitive s k env = let (>>) = Inner.seq in Inner.get >>= fun step -> try - let (tacs,defs) = list_of_sensitive s k env step in - Inner.set { step with solution = defs } >> + let (tacs,defs,effs) = list_of_sensitive s k env step in + Inner.set { step with solution = defs; + side_effects = Declareops.union_side_effects step.side_effects + (Declareops.flatten_side_effects effs) } >> Inner.return tacs with e when catchable_exception e -> tclZERO e env @@ -388,19 +400,22 @@ let tclGOALBINDU s k = (* No backtracking can happen here, hence, as opposed to the dispatch tacticals, everything is done in one step. *) let sensitive_on_proofview s env step = - let wrap g ((defs, partial_list) as partial_res) = + let wrap g ((defs, partial_list, partial_eff) as partial_res) = match Goal.advance defs g with - | None ->partial_res + | None -> partial_res | Some g -> - let {Goal.subgoals = sg } , d' = Goal.eval s env defs g in - (d',sg::partial_list) + let { Goal.subgoals = sg } , d', eff = Goal.eval s env defs g in + (d', sg::partial_list, eff::partial_eff) in - let ( new_defs , combed_subgoals ) = - List.fold_right wrap step.comb (step.solution,[]) + let ( new_defs , combed_subgoals, side_effects ) = + List.fold_right wrap step.comb (step.solution,[],[]) in { step with solution = new_defs; - comb = List.flatten combed_subgoals } + comb = List.flatten combed_subgoals; + side_effects = + Declareops.union_side_effects + (Declareops.flatten_side_effects side_effects) step.side_effects } let tclSENSITIVE s env = (* spiwack: convenience notations, waiting for ocaml 3.12 *) let (>>=) = Inner.bind in @@ -440,18 +455,26 @@ module V82 = struct Inner.get >>= fun ps -> try let tac evd gl = - let glsigma = tac { Evd.it = gl ; Evd.sigma = evd } in + let glsigma = + tac { Evd.it = gl ; sigma = evd; eff = Declareops.no_seff } in let sigma = glsigma.Evd.sigma in let g = glsigma.Evd.it in - ( g , sigma ) + ( g, sigma, glsigma.Evd.eff ) in (* Old style tactics expect the goals normalized with respect to evars. *) - let (initgoals,initevd) = - Goal.list_map Goal.V82.nf_evar ps.comb ps.solution + let no_side_eff f x y = + let a, b = f x y in a, b, Declareops.no_seff in + let (initgoals,initevd, eff1) = + Goal.list_map (no_side_eff Goal.V82.nf_evar) ps.comb ps.solution in - let (goalss,evd) = Goal.list_map tac initgoals initevd in + let (goalss,evd,eff2) = Goal.list_map tac initgoals initevd in let sgs = List.flatten goalss in - Inner.set { ps with solution=evd ; comb=sgs } + Inner.set { ps with solution=evd ; comb=sgs; + side_effects = + Declareops.union_side_effects ps.side_effects + (Declareops.union_side_effects + (Declareops.flatten_side_effects eff1) + (Declareops.flatten_side_effects eff2)) } with e when catchable_exception e -> tclZERO e env @@ -471,12 +494,12 @@ module V82 = struct (* Returns the open goals of the proofview together with the evar_map to interprete them. *) - let goals { comb = comb ; solution = solution } = - { Evd.it = comb ; sigma = solution} + let goals { comb = comb ; solution = solution; side_effects=eff } = + { Evd.it = comb ; sigma = solution; eff=eff} - let top_goals { initial=initial ; solution=solution } = + let top_goals { initial=initial ; solution=solution; side_effects=eff } = let goals = List.map (fun (t,_) -> Goal.V82.build (fst (Term.destEvar t))) initial in - { Evd.it = goals ; sigma=solution } + { Evd.it = goals ; sigma=solution; eff=eff } let top_evars { initial=initial } = let evars_of_initial (c,_) = diff --git a/proofs/proofview.mli b/proofs/proofview.mli index ff327ab3b..2869e75e1 100644 --- a/proofs/proofview.mli +++ b/proofs/proofview.mli @@ -44,8 +44,11 @@ val init : (Environ.env * Term.types) list -> proofview val finished : proofview -> bool (* Returns the current value of the proofview partial proofs. *) -val return : proofview -> (constr*types) list +val return : proofview -> constr -> constr * Declareops.side_effects +val partial_proof : proofview -> constr list +val initial_goals : proofview -> (constr * types) list +val emit_side_effects : Declareops.side_effects -> proofview -> proofview (*** Focusing operations ***) @@ -181,10 +184,6 @@ val tclSENSITIVE : Goal.subgoals Goal.sensitive -> unit tactic val in_proofview : proofview -> (Evd.evar_map -> 'a) -> 'a - - - - (* Notations for building tactics. *) module Notations : sig (* Goal.bind *) diff --git a/proofs/refiner.ml b/proofs/refiner.ml index dfc9c0a63..06e35284d 100644 --- a/proofs/refiner.ml +++ b/proofs/refiner.ml @@ -16,6 +16,7 @@ open Logic let sig_it x = x.it +let sig_eff x = x.eff let project x = x.sigma (* Getting env *) @@ -25,7 +26,7 @@ let pf_hyps gls = named_context_of_val (Goal.V82.hyps (project gls) (sig_it gls) let refiner pr goal_sigma = let (sgl,sigma') = prim_refiner pr goal_sigma.sigma goal_sigma.it in - {it=sgl; sigma = sigma'} + { it = sgl; sigma = sigma'; eff = goal_sigma.eff } let norm_evar_tac gl = refiner (Change_evars) gl @@ -34,18 +35,19 @@ let norm_evar_tac gl = refiner (Change_evars) gl (*********************) -let unpackage glsig = (ref (glsig.sigma)),glsig.it +let unpackage glsig = (ref (glsig.sigma)), glsig.it, (ref (glsig.eff)) -let repackage r v = {it=v;sigma = !r} +let repackage e r v = {it = v; sigma = !r; eff = !e} -let apply_sig_tac r tac g = +let apply_sig_tac eff r tac g = check_for_interrupt (); (* Breakpoint *) - let glsigma = tac (repackage r g) in + let glsigma = tac (repackage eff r g) in r := glsigma.sigma; + eff := glsigma.eff; glsigma.it (* [goal_goal_list : goal sigma -> goal list sigma] *) -let goal_goal_list gls = {it=[gls.it];sigma=gls.sigma} +let goal_goal_list gls = {it=[gls.it]; sigma=gls.sigma; eff=gls.eff} (* forces propagation of evar constraints *) let tclNORMEVAR = norm_evar_tac @@ -69,35 +71,39 @@ let tclFAIL lvl s g = raise (FailError (lvl,lazy s)) let tclFAIL_lazy lvl s g = raise (FailError (lvl,s)) let start_tac gls = - let (sigr,g) = unpackage gls in - (sigr,[g]) + let sigr, g, eff = unpackage gls in + (sigr, [g], eff) -let finish_tac (sigr,gl) = repackage sigr gl +let finish_tac (sigr,gl,eff) = repackage eff sigr gl (* Apply [tacfi.(i)] on the first n subgoals, [tacli.(i)] on the last m subgoals, and [tac] on the others *) -let thens3parts_tac tacfi tac tacli (sigr,gs) = +let thens3parts_tac tacfi tac tacli (sigr,gs,eff) = let nf = Array.length tacfi in let nl = Array.length tacli in let ng = List.length gs in if ng - apply_sig_tac sigr (if i=ng-nl then tacli.(nl-ng+i) else tac)) + apply_sig_tac eff sigr (if i=ng-nl then tacli.(nl-ng+i) else tac)) 0 gs) in - (sigr,List.flatten gll) + (sigr,List.flatten gll, eff) (* Apply [taci.(i)] on the first n subgoals and [tac] on the others *) let thensf_tac taci tac = thens3parts_tac taci tac [||] (* Apply [tac i] on the ith subgoal (no subgoals number check) *) -let thensi_tac tac (sigr,gs) = +let thensi_tac tac (sigr,gs,eff) = let gll = - List.map_i (fun i -> apply_sig_tac sigr (tac i)) 1 gs in - (sigr, List.flatten gll) + List.map_i (fun i -> apply_sig_tac eff sigr (tac i)) 1 gs in + (sigr, List.flatten gll, eff) let then_tac tac = thensf_tac [||] tac +let non_existent_goal n = + errorlabstrm ("No such goal: "^(string_of_int n)) + (str"Trying to apply a tactic to a non existent goal") + (* [tclTHENS3PARTS tac1 [|t1 ; ... ; tn|] tac2 [|t'1 ; ... ; t'm|] gls] applies the tactic [tac1] to [gls] then, applies [t1], ..., [tn] to the first [n] resulting subgoals, [t'1], ..., [t'm] to the last [m] @@ -192,9 +198,9 @@ let tclSHOWHYPS (tac : tactic) (goal: Goal.goal Evd.sigma) :Proof_type.goal list Evd.sigma = let oldhyps:Context.named_context = pf_hyps goal in let rslt:Proof_type.goal list Evd.sigma = tac goal in - let {it=gls;sigma=sigma} = rslt in + let {it=gls;sigma=sigma;eff=eff} = rslt in let hyps:Context.named_context list = - List.map (fun gl -> pf_hyps { it = gl; sigma=sigma}) gls in + List.map (fun gl -> pf_hyps { it = gl; sigma=sigma; eff=eff}) gls in let newhyps = List.map (fun hypl -> List.subtract hypl oldhyps) hyps in let emacs_str s = @@ -244,8 +250,8 @@ let tclORELSE_THEN t1 t2then t2else gls = with | None -> t2else gls | Some sgl -> - let (sigr,gl) = unpackage sgl in - finish_tac (then_tac t2then (sigr,gl)) + let sigr, gl, eff = unpackage sgl in + finish_tac (then_tac t2then (sigr,gl,eff)) (* TRY f tries to apply f, and if it fails, leave the goal unchanged *) let tclTRY f = (tclORELSE0 f tclIDTAC) @@ -350,23 +356,24 @@ let tclIDTAC_list gls = gls (* first_goal : goal list sigma -> goal sigma *) let first_goal gls = - let gl = gls.it and sig_0 = gls.sigma in + let gl = gls.it and sig_0 = gls.sigma and eff_0 = gls.eff in if List.is_empty gl then error "first_goal"; - { it = List.hd gl; sigma = sig_0 } + { it = List.hd gl; sigma = sig_0 ; eff = eff_0 } (* goal_goal_list : goal sigma -> goal list sigma *) let goal_goal_list gls = - let gl = gls.it and sig_0 = gls.sigma in { it = [gl]; sigma = sig_0 } + let gl = gls.it and sig_0 = gls.sigma and eff_0 = gls.eff in + { it = [gl]; sigma = sig_0; eff = eff_0 } (* tactic -> tactic_list : Apply a tactic to the first goal in the list *) let apply_tac_list tac glls = - let (sigr,lg) = unpackage glls in + let (sigr,lg,eff) = unpackage glls in match lg with | (g1::rest) -> - let gl = apply_sig_tac sigr tac g1 in - repackage sigr (gl@rest) + let gl = apply_sig_tac eff sigr tac g1 in + repackage eff sigr (gl@rest) | _ -> error "apply_tac_list" let then_tactic_list tacl1 tacl2 glls = diff --git a/proofs/refiner.mli b/proofs/refiner.mli index c198958e3..b2e753ea3 100644 --- a/proofs/refiner.mli +++ b/proofs/refiner.mli @@ -16,15 +16,17 @@ open Logic (** The refiner (handles primitive rules and high-level tactics). *) val sig_it : 'a sigma -> 'a +val sig_eff : 'a sigma -> Declareops.side_effects val project : 'a sigma -> evar_map val pf_env : goal sigma -> Environ.env val pf_hyps : goal sigma -> named_context -val unpackage : 'a sigma -> evar_map ref * 'a -val repackage : evar_map ref -> 'a -> 'a sigma +val unpackage : 'a sigma -> evar_map ref * 'a * Declareops.side_effects ref +val repackage : Declareops.side_effects ref -> evar_map ref -> 'a -> 'a sigma val apply_sig_tac : - evar_map ref -> (goal sigma -> goal list sigma) -> goal -> goal list + Declareops.side_effects ref -> evar_map ref -> + (goal sigma -> goal list sigma) -> goal -> goal list val refiner : rule -> tactic diff --git a/proofs/tacmach.ml b/proofs/tacmach.ml index 813a0d850..fad46c99b 100644 --- a/proofs/tacmach.ml +++ b/proofs/tacmach.ml @@ -22,7 +22,7 @@ open Proof_type open Logic open Refiner -let re_sig it gc = { it = it; sigma = gc } +let re_sig it eff gc = { it = it; sigma = gc; eff = eff } (**************************************************************) (* Operations for handling terms under a local typing context *) @@ -36,6 +36,7 @@ let repackage = Refiner.repackage let apply_sig_tac = Refiner.apply_sig_tac let sig_it = Refiner.sig_it +let sig_eff = Refiner.sig_eff let project = Refiner.project let pf_env = Refiner.pf_env let pf_hyps = Refiner.pf_hyps diff --git a/proofs/tacmach.mli b/proofs/tacmach.mli index 1c0ab2d14..19eb9ba7c 100644 --- a/proofs/tacmach.mli +++ b/proofs/tacmach.mli @@ -27,14 +27,16 @@ type 'a sigma = 'a Evd.sigma;; type tactic = Proof_type.tactic;; val sig_it : 'a sigma -> 'a +val sig_eff : 'a sigma -> Declareops.side_effects val project : goal sigma -> evar_map -val re_sig : 'a -> evar_map -> 'a sigma +val re_sig : 'a -> Declareops.side_effects -> evar_map -> 'a sigma -val unpackage : 'a sigma -> evar_map ref * 'a -val repackage : evar_map ref -> 'a -> 'a sigma +val unpackage : 'a sigma -> evar_map ref * 'a * Declareops.side_effects ref +val repackage : Declareops.side_effects ref -> evar_map ref -> 'a -> 'a sigma val apply_sig_tac : - evar_map ref -> (goal sigma -> (goal list) sigma) -> goal -> (goal list) + Declareops.side_effects ref -> evar_map ref -> + (goal sigma -> (goal list) sigma) -> goal -> (goal list) val pf_concl : goal sigma -> types val pf_env : goal sigma -> env diff --git a/tactics/auto.ml b/tactics/auto.ml index fab324962..29eb9e9ec 100644 --- a/tactics/auto.ml +++ b/tactics/auto.ml @@ -983,7 +983,7 @@ let pr_applicable_hint () = match glss.Evd.it with | [] -> Errors.error "No focused goal." | g::_ -> - let gl = { Evd.it = g; sigma = glss.Evd.sigma } in + let gl = { Evd.it = g; sigma = glss.Evd.sigma; eff = Declareops.no_seff } in pr_hint_term (pf_concl gl) (* displays the whole hint database db *) diff --git a/tactics/class_tactics.ml4 b/tactics/class_tactics.ml4 index bfe80d708..21678c971 100644 --- a/tactics/class_tactics.ml4 +++ b/tactics/class_tactics.ml4 @@ -291,15 +291,16 @@ let make_autogoal_hints = unfreeze (Some (only_classes, sign, hints)); hints let lift_tactic tac (f : goal list sigma -> autoinfo -> autogoal list sigma) : 'a tac = - { skft = fun sk fk {it = gl,hints; sigma=s} -> - let res = try Some (tac {it=gl; sigma=s}) with e when catchable e -> None in + { skft = fun sk fk {it = gl,hints; sigma=s;eff=eff} -> + let res = try Some (tac {it=gl; sigma=s;eff=eff}) + with e when catchable e -> None in match res with | Some gls -> sk (f gls hints) fk | None -> fk () } let intro_tac : atac = lift_tactic Tactics.intro - (fun {it = gls; sigma = s} info -> + (fun {it = gls; sigma = s;eff=eff} info -> let gls' = List.map (fun g' -> let env = Goal.V82.env s g' in @@ -308,13 +309,13 @@ let intro_tac : atac = (true,false,false) info.only_classes None (List.hd context) in let ldb = Hint_db.add_list hint info.hints in (g', { info with is_evar = None; hints = ldb; auto_last_tac = lazy (str"intro") })) gls - in {it = gls'; sigma = s}) + in {it = gls'; sigma = s;eff=eff}) let normevars_tac : atac = - { skft = fun sk fk {it = (gl, info); sigma = s} -> + { skft = fun sk fk {it = (gl, info); sigma = s; eff=eff} -> let gl', sigma' = Goal.V82.nf_evar s gl in let info' = { info with auto_last_tac = lazy (str"normevars") } in - sk {it = [gl', info']; sigma = sigma'} fk } + sk {it = [gl', info']; sigma = sigma';eff=eff} fk } (* Ordering of states is lexicographic on the number of remaining goals. *) let compare (pri, _, _, res) (pri', _, _, res') = @@ -329,9 +330,9 @@ let or_tac (x : 'a tac) (y : 'a tac) : 'a tac = { skft = fun sk fk gls -> x.skft sk (fun () -> y.skft sk fk gls) gls } let hints_tac hints = - { skft = fun sk fk {it = gl,info; sigma = s} -> + { skft = fun sk fk {it = gl,info; sigma = s;eff=eff} -> let concl = Goal.V82.concl s gl in - let tacgl = {it = gl; sigma = s} in + let tacgl = {it = gl; sigma = s;eff=eff} in let poss = e_possible_resolve hints info.hints concl in let rec aux i foundone = function | (tac, _, b, name, pp) :: tl -> @@ -343,7 +344,7 @@ let hints_tac hints = in (match res with | None -> aux i foundone tl - | Some {it = gls; sigma = s'} -> + | Some {it = gls; sigma = s';eff=eff} -> if !typeclasses_debug then msg_debug (pr_depth (i :: info.auto_depth) ++ str": " ++ Lazy.force pp ++ str" on" ++ spc () ++ pr_ev s gl); @@ -375,11 +376,11 @@ let hints_tac hints = hints = if b && not (Environ.eq_named_context_val (Goal.V82.hyps s' g) (Goal.V82.hyps s' gl)) then make_autogoal_hints info.only_classes - ~st:(Hint_db.transparent_state info.hints) {it = g; sigma = s'} + ~st:(Hint_db.transparent_state info.hints) {it = g; sigma = s';eff=eff} else info.hints; auto_cut = derivs } in g, info) 1 newgls in - let glsv = {it = gls'; sigma = s'} in + let glsv = {it = gls'; sigma = s';eff=eff} in sk glsv fk) | [] -> if not foundone && !typeclasses_debug then @@ -422,14 +423,14 @@ let then_list (second : atac) (sk : (auto_result, 'a) sk) : (auto_result, 'a) sk fk) else fk' in aux s' (gls'::acc) fk'' gls) - fk {it = (gl,info); sigma = s}) + fk {it = (gl,info); sigma = s; eff = Declareops.no_seff }) | [] -> Somek2 (List.rev acc, s, fk) - in fun {it = gls; sigma = s} fk -> + in fun {it = gls; sigma = s; eff = eff} fk -> let rec aux' = function | Nonek2 -> fk () | Somek2 (res, s', fk') -> let goals' = List.concat res in - sk {it = goals'; sigma = s'} (fun () -> aux' (fk' ())) + sk {it = goals'; sigma = s'; eff = eff } (fun () -> aux' (fk' ())) in aux' (aux s [] (fun () -> Nonek2) gls) let then_tac (first : atac) (second : atac) : atac = @@ -468,8 +469,8 @@ let cut_of_hints h = let make_autogoals ?(only_classes=true) ?(st=full_transparent_state) hints gs evm' = let cut = cut_of_hints hints in { it = List.map_i (fun i g -> - let (gl, auto) = make_autogoal ~only_classes ~st cut (Some (fst g)) {it = snd g; sigma = evm'} in - (gl, { auto with auto_depth = [i]})) 1 gs; sigma = evm' } + let (gl, auto) = make_autogoal ~only_classes ~st cut (Some (fst g)) {it = snd g; sigma = evm'; eff = Declareops.no_seff } in + (gl, { auto with auto_depth = [i]})) 1 gs; sigma = evm'; eff = Declareops.no_seff } let get_result r = match r with @@ -494,11 +495,11 @@ let eauto_tac ?limit hints = | Some limit -> fix_limit limit (eauto_tac hints) let eauto ?(only_classes=true) ?st ?limit hints g = - let gl = { it = make_autogoal ~only_classes ?st (cut_of_hints hints) None g; sigma = project g } in + let gl = { it = make_autogoal ~only_classes ?st (cut_of_hints hints) None g; sigma = project g; eff=g.eff } in match run_tac (eauto_tac ?limit hints) gl with | None -> raise Not_found - | Some {it = goals; sigma = s} -> - {it = List.map fst goals; sigma = s} + | Some {it = goals; sigma = s; eff=eff} -> + {it = List.map fst goals; sigma = s; eff=eff} let real_eauto st ?limit hints p evd = let rec aux evd fails = @@ -544,7 +545,7 @@ let resolve_one_typeclass env ?(sigma=Evd.empty) gl = let nc, gl, subst, _ = Evarutil.push_rel_context_to_named_context env gl in let (gl,t,sigma) = Goal.V82.mk_goal sigma nc gl Store.empty in - let gls = { it = gl ; sigma = sigma } in + let gls = { it = gl ; sigma = sigma; eff= Declareops.no_seff } in (* XXX eff *) let hints = searchtable_map typeclasses_db in let gls' = eauto ?limit:!typeclasses_depth ~st:(Hint_db.transparent_state hints) [hints] gls in let evd = sig_sig gls' in diff --git a/tactics/elimschemes.ml b/tactics/elimschemes.ml index 62d13c0a6..9c020930c 100644 --- a/tactics/elimschemes.ml +++ b/tactics/elimschemes.ml @@ -27,7 +27,7 @@ let optimize_non_type_induction_scheme kind dep sort ind = (* in case the inductive has a type elimination, generates only one induction scheme, the other ones share the same code with the apropriate type *) - let cte = find_scheme kind ind in + let cte, eff = find_scheme kind ind in let c = mkConst cte in let t = type_of_constant (Global.env()) cte in let (mib,mip) = Global.lookup_inductive ind in @@ -39,24 +39,24 @@ let optimize_non_type_induction_scheme kind dep sort ind = mib.mind_nparams_rec else mib.mind_nparams in - snd (weaken_sort_scheme (new_sort_in_family sort) npars c t) + snd (weaken_sort_scheme (new_sort_in_family sort) npars c t), eff else - build_induction_scheme (Global.env()) Evd.empty ind dep sort + build_induction_scheme (Global.env()) Evd.empty ind dep sort, Declareops.no_seff let build_induction_scheme_in_type dep sort ind = build_induction_scheme (Global.env()) Evd.empty ind dep sort let rect_scheme_kind_from_type = declare_individual_scheme_object "_rect_nodep" - (build_induction_scheme_in_type false InType) + (fun x -> build_induction_scheme_in_type false InType x, Declareops.no_seff) let rect_scheme_kind_from_prop = declare_individual_scheme_object "_rect" ~aux:"_rect_from_prop" - (build_induction_scheme_in_type false InType) + (fun x -> build_induction_scheme_in_type false InType x, Declareops.no_seff) let rect_dep_scheme_kind_from_type = declare_individual_scheme_object "_rect" ~aux:"_rect_from_type" - (build_induction_scheme_in_type true InType) + (fun x -> build_induction_scheme_in_type true InType x, Declareops.no_seff) let ind_scheme_kind_from_type = declare_individual_scheme_object "_ind_nodep" @@ -85,24 +85,24 @@ let build_case_analysis_scheme_in_type dep sort ind = let case_scheme_kind_from_type = declare_individual_scheme_object "_case_nodep" - (build_case_analysis_scheme_in_type false InType) + (fun x -> build_case_analysis_scheme_in_type false InType x, Declareops.no_seff) let case_scheme_kind_from_prop = declare_individual_scheme_object "_case" ~aux:"_case_from_prop" - (build_case_analysis_scheme_in_type false InType) + (fun x -> build_case_analysis_scheme_in_type false InType x, Declareops.no_seff) let case_dep_scheme_kind_from_type = declare_individual_scheme_object "_case" ~aux:"_case_from_type" - (build_case_analysis_scheme_in_type true InType) + (fun x -> build_case_analysis_scheme_in_type true InType x, Declareops.no_seff) let case_dep_scheme_kind_from_type_in_prop = declare_individual_scheme_object "_casep_dep" - (build_case_analysis_scheme_in_type true InProp) + (fun x -> build_case_analysis_scheme_in_type true InProp x, Declareops.no_seff) let case_dep_scheme_kind_from_prop = declare_individual_scheme_object "_case_dep" - (build_case_analysis_scheme_in_type true InType) + (fun x -> build_case_analysis_scheme_in_type true InType x, Declareops.no_seff) let case_dep_scheme_kind_from_prop_in_prop = declare_individual_scheme_object "_casep" - (build_case_analysis_scheme_in_type true InProp) + (fun x -> build_case_analysis_scheme_in_type true InProp x, Declareops.no_seff) diff --git a/tactics/eqschemes.ml b/tactics/eqschemes.ml index cc1162168..7aac37d1b 100644 --- a/tactics/eqschemes.ml +++ b/tactics/eqschemes.ml @@ -180,7 +180,7 @@ let build_sym_scheme env ind = let sym_scheme_kind = declare_individual_scheme_object "_sym_internal" - (fun ind -> build_sym_scheme (Global.env() (* side-effect! *)) ind) + (fun ind -> build_sym_scheme (Global.env() (* side-effect! *)) ind, Declareops.no_seff) (**********************************************************************) (* Build the involutivity of symmetry for an inductive type *) @@ -201,7 +201,8 @@ let sym_scheme_kind = let build_sym_involutive_scheme env ind = let (mib,mip as specif),nrealargs,realsign,paramsctxt,paramsctxt1 = get_sym_eq_data env ind in - let sym = mkConst (find_scheme sym_scheme_kind ind) in + let c, eff = find_scheme sym_scheme_kind ind in + let sym = mkConst c in let (eq,eqrefl) = get_coq_eq () in let cstr n = mkApp (mkConstruct(ind,1),extended_rel_vect n paramsctxt) in let varH = fresh env (default_id_of_sort (snd (mind_arity mip))) in @@ -236,7 +237,8 @@ let build_sym_involutive_scheme env ind = [|mkRel 1|]])|]]); mkRel 1|])), mkRel 1 (* varH *), - [|mkApp(eqrefl,[|applied_ind_C;cstr (nrealargs+1)|])|])))) + [|mkApp(eqrefl,[|applied_ind_C;cstr (nrealargs+1)|])|])))), + eff let sym_involutive_scheme_kind = declare_individual_scheme_object "_sym_involutive" @@ -305,8 +307,10 @@ let sym_involutive_scheme_kind = let build_l2r_rew_scheme dep env ind kind = let (mib,mip as specif),nrealargs,realsign,paramsctxt,paramsctxt1 = get_sym_eq_data env ind in - let sym = mkConst (find_scheme sym_scheme_kind ind) in - let sym_involutive = mkConst (find_scheme sym_involutive_scheme_kind ind) in + let c, eff = find_scheme sym_scheme_kind ind in + let sym = mkConst c in + let c, eff' = find_scheme sym_involutive_scheme_kind ind in + let sym_involutive = mkConst c in let (eq,eqrefl) = get_coq_eq () in let cstr n p = mkApp (mkConstruct(ind,1), @@ -384,7 +388,8 @@ let build_l2r_rew_scheme dep env ind kind = Array.append (extended_rel_vect 3 mip.mind_arity_ctxt) [|mkVar varH|]), [|main_body|]) else - main_body)))))) + main_body)))))), + Declareops.union_side_effects eff' eff (**********************************************************************) (* Build the left-to-right rewriting lemma for hypotheses associated *) @@ -613,7 +618,7 @@ let rew_l2r_dep_scheme_kind = (**********************************************************************) let rew_r2l_dep_scheme_kind = declare_individual_scheme_object "_rew_dep" - (fun ind -> build_r2l_rew_scheme true (Global.env()) ind InType) + (fun ind -> build_r2l_rew_scheme true (Global.env()) ind InType,Declareops.no_seff) (**********************************************************************) (* Dependent rewrite from right-to-left in hypotheses *) @@ -623,7 +628,7 @@ let rew_r2l_dep_scheme_kind = (**********************************************************************) let rew_r2l_forward_dep_scheme_kind = declare_individual_scheme_object "_rew_fwd_dep" - (fun ind -> build_r2l_forward_rew_scheme true (Global.env()) ind InType) + (fun ind -> build_r2l_forward_rew_scheme true (Global.env()) ind InType,Declareops.no_seff) (**********************************************************************) (* Dependent rewrite from left-to-right in hypotheses *) @@ -633,7 +638,7 @@ let rew_r2l_forward_dep_scheme_kind = (**********************************************************************) let rew_l2r_forward_dep_scheme_kind = declare_individual_scheme_object "_rew_fwd_r_dep" - (fun ind -> build_l2r_forward_rew_scheme true (Global.env()) ind InType) + (fun ind -> build_l2r_forward_rew_scheme true (Global.env()) ind InType,Declareops.no_seff) (**********************************************************************) (* Non-dependent rewrite from either left-to-right in conclusion or *) @@ -647,7 +652,7 @@ let rew_l2r_forward_dep_scheme_kind = let rew_l2r_scheme_kind = declare_individual_scheme_object "_rew_r" (fun ind -> fix_r2l_forward_rew_scheme - (build_r2l_forward_rew_scheme false (Global.env()) ind InType)) + (build_r2l_forward_rew_scheme false (Global.env()) ind InType), Declareops.no_seff) (**********************************************************************) (* Non-dependent rewrite from either right-to-left in conclusion or *) @@ -657,7 +662,7 @@ let rew_l2r_scheme_kind = (**********************************************************************) let rew_r2l_scheme_kind = declare_individual_scheme_object "_rew" - (fun ind -> build_r2l_rew_scheme false (Global.env()) ind InType) + (fun ind -> build_r2l_rew_scheme false (Global.env()) ind InType, Declareops.no_seff) (* End of rewriting schemes *) @@ -728,4 +733,4 @@ let build_congr env (eq,refl) ind = let congr_scheme_kind = declare_individual_scheme_object "_congr" (fun ind -> (* May fail if equality is not defined *) - build_congr (Global.env()) (get_coq_eq ()) ind) + build_congr (Global.env()) (get_coq_eq ()) ind, Declareops.no_seff) diff --git a/tactics/eqschemes.mli b/tactics/eqschemes.mli index 31a96e6dc..72412d12d 100644 --- a/tactics/eqschemes.mli +++ b/tactics/eqschemes.mli @@ -23,7 +23,8 @@ val rew_r2l_dep_scheme_kind : individual scheme_kind val rew_r2l_scheme_kind : individual scheme_kind val build_r2l_rew_scheme : bool -> env -> inductive -> sorts_family -> constr -val build_l2r_rew_scheme : bool -> env -> inductive -> sorts_family -> constr +val build_l2r_rew_scheme : + bool -> env -> inductive -> sorts_family -> constr * Declareops.side_effects val build_r2l_forward_rew_scheme : bool -> env -> inductive -> sorts_family -> constr val build_l2r_forward_rew_scheme : @@ -34,7 +35,8 @@ val build_l2r_forward_rew_scheme : val build_sym_scheme : env -> inductive -> constr val sym_scheme_kind : individual scheme_kind -val build_sym_involutive_scheme : env -> inductive -> constr +val build_sym_involutive_scheme : + env -> inductive -> constr * Declareops.side_effects val sym_involutive_scheme_kind : individual scheme_kind (** Builds a congruence scheme for an equality type *) diff --git a/tactics/equality.ml b/tactics/equality.ml index b991defe7..f5ab039b4 100644 --- a/tactics/equality.ml +++ b/tactics/equality.ml @@ -283,12 +283,12 @@ let find_elim hdcncl lft2rgt dep cls ot gl = begin try let _ = Global.lookup_constant c1' in - mkConst c1' + mkConst c1', gl with Not_found -> let rwr_thm = Label.to_string l' in error ("Cannot find rewrite principle "^rwr_thm^".") end - | _ -> pr1 + | _ -> pr1, gl end | _ -> (* cannot occur since we checked that we are in presence of @@ -308,7 +308,10 @@ let find_elim hdcncl lft2rgt dep cls ot gl = | true, _, false -> rew_r2l_forward_dep_scheme_kind in match kind_of_term hdcncl with - | Ind ind -> mkConst (find_scheme scheme_name ind) + | Ind ind -> + let c, eff = find_scheme scheme_name ind in + let gl = {gl with eff = Declareops.union_side_effects eff gl.eff } in + mkConst c, gl | _ -> assert false let type_of_clause gl = function @@ -319,7 +322,7 @@ let leibniz_rewrite_ebindings_clause cls lft2rgt tac sigma c t l with_evars frze let isatomic = isProd (whd_zeta hdcncl) in let dep_fun = if isatomic then dependent else dependent_no_evar in let dep = dep_proof_ok && dep_fun c (type_of_clause gl cls) in - let elim = find_elim hdcncl lft2rgt dep cls (Some t) gl in + let elim, gl = find_elim hdcncl lft2rgt dep cls (Some t) gl in general_elim_clause with_evars frzevars tac cls sigma c t l (match lft2rgt with None -> false | Some b -> b) {elimindex = None; elimbody = (elim,NoBindings)} gl @@ -775,14 +778,16 @@ let ind_scheme_of_eq lbeq = let kind = if kind == InProp then Elimschemes.ind_scheme_kind_from_prop else Elimschemes.ind_scheme_kind_from_type in - mkConst (find_scheme kind (destInd lbeq.eq)) + let c, eff = find_scheme kind (destInd lbeq.eq) in + mkConst c, eff let discrimination_pf e (t,t1,t2) discriminator lbeq = - let i = build_coq_I () in - let absurd_term = build_coq_False () in - let eq_elim = ind_scheme_of_eq lbeq in - (applist (eq_elim, [t;t1;mkNamedLambda e t discriminator;i;t2]), absurd_term) + let i = build_coq_I () in + let absurd_term = build_coq_False () in + let eq_elim, eff = ind_scheme_of_eq lbeq in + (applist (eq_elim, [t;t1;mkNamedLambda e t discriminator;i;t2]), absurd_term), + eff let eq_baseid = Id.of_string "e" @@ -795,17 +800,19 @@ let apply_on_clause (f,t) clause = | _ -> errorlabstrm "" (str "Ill-formed clause applicator.")) in clenv_fchain argmv f_clause clause -let discr_positions env sigma (lbeq,eqn,(t,t1,t2)) eq_clause cpath dirn sort = +let discr_positions env sigma (lbeq,eqn,(t,t1,t2)) eq_clause cpath dirn sort gl= let e = next_ident_away eq_baseid (ids_of_context env) in let e_env = push_named (e,None,t) env in let discriminator = build_discriminator sigma e_env dirn (mkVar e) sort cpath in - let (pf, absurd_term) = discrimination_pf e (t,t1,t2) discriminator lbeq in + let (pf, absurd_term), eff = + discrimination_pf e (t,t1,t2) discriminator lbeq in let pf_ty = mkArrow eqn absurd_term in let absurd_clause = apply_on_clause (pf,pf_ty) eq_clause in let pf = clenv_value_cast_meta absurd_clause in + let gl = {gl with eff = Declareops.union_side_effects eff gl.eff } in tclTHENS (cut_intro absurd_term) - [onLastHypId gen_absurdity; refine pf] + [onLastHypId gen_absurdity; refine pf] gl let discrEq (lbeq,_,(t,t1,t2) as u) eq_clause gls = let sigma = eq_clause.evd in @@ -1134,7 +1141,7 @@ let set_eq_dec_scheme_kind k = eq_dec_scheme_kind_name := (fun _ -> k) let eqdep_dec = qualid_of_string "Coq.Logic.Eqdep_dec" -let inject_if_homogenous_dependent_pair env sigma (eq,_,(t,t1,t2)) = +let inject_if_homogenous_dependent_pair env sigma (eq,_,(t,t1,t2)) gl = (* fetch the informations of the pair *) let ceq = constr_of_global Coqlib.glob_eq in let sigTconstr () = (Coqlib.build_sigma_type()).Coqlib.typ in @@ -1153,19 +1160,19 @@ let inject_if_homogenous_dependent_pair env sigma (eq,_,(t,t1,t2)) = Library.require_library [Loc.ghost,eqdep_dec] (Some false); let inj2 = Coqlib.coq_constant "inj_pair2_eq_dec is missing" ["Logic";"Eqdep_dec"] "inj_pair2_eq_dec" in + let c, eff = find_scheme (!eq_dec_scheme_kind_name()) ind in + let gl = { gl with eff = Declareops.union_side_effects eff gl.eff } in (* cut with the good equality and prove the requested goal *) tclTHENS (cut (mkApp (ceq,new_eq_args))) [tclIDTAC; tclTHEN (apply ( - mkApp(inj2, - [|ar1.(0);mkConst (find_scheme (!eq_dec_scheme_kind_name()) ind); - ar1.(1);ar1.(2);ar1.(3);ar2.(3)|]) + mkApp(inj2,[|ar1.(0);mkConst c;ar1.(1);ar1.(2);ar1.(3);ar2.(3)|]) )) (Auto.trivial [] []) - ] + ] gl (* not a dep eq or no decidable type found *) end else raise Not_dep_pair -let injEqThen tac l2r (eq,_,(t,t1,t2) as u) eq_clause = +let injEqThen tac l2r (eq,_,(t,t1,t2) as u) eq_clause gl = let sigma = eq_clause.evd in let env = eq_clause.env in match find_positions env sigma t1 t2 with @@ -1178,10 +1185,10 @@ let injEqThen tac l2r (eq,_,(t,t1,t2) as u) eq_clause = | Inr [([],_,_)] when Flags.version_strictly_greater Flags.V8_3 -> errorlabstrm "Equality.inj" (str"Nothing to inject.") | Inr posns -> - try inject_if_homogenous_dependent_pair env sigma u + try inject_if_homogenous_dependent_pair env sigma u gl with e when (Errors.noncritical e || e = Not_dep_pair) -> inject_at_positions env sigma l2r u eq_clause posns - (tac (clenv_value eq_clause)) + (tac (clenv_value eq_clause)) gl let postInjEqTac ipats c n = match ipats with @@ -1255,7 +1262,7 @@ let swapEquandsInConcl gls = let bareRevSubstInConcl lbeq body (t,e1,e2) gls = (* find substitution scheme *) - let eq_elim = find_elim lbeq.eq (Some false) false None None gls in + let eq_elim, gls = find_elim lbeq.eq (Some false) false None None gls in (* build substitution predicate *) let p = lambda_create (pf_env gls) (t,body) in (* apply substitution scheme *) diff --git a/tactics/extratactics.ml4 b/tactics/extratactics.ml4 index e14c23421..ae617a4d8 100644 --- a/tactics/extratactics.ml4 +++ b/tactics/extratactics.ml4 @@ -115,7 +115,8 @@ TACTIC EXTEND ediscriminate END let discrHyp id gl = - discr_main {it = Term.mkVar id,NoBindings; sigma = Refiner.project gl} gl + discr_main {it = Term.mkVar id,NoBindings; sigma = Refiner.project gl; + eff = gl.eff} gl let injection_main c = elimOnConstrWithHoles (injClause None) false c @@ -159,7 +160,8 @@ TACTIC EXTEND einjection_as END let injHyp id gl = - injection_main { it = Term.mkVar id,NoBindings; sigma = Refiner.project gl } gl + injection_main { it = Term.mkVar id,NoBindings; sigma = Refiner.project gl; + eff = gl.eff } gl TACTIC EXTEND dependent_rewrite | [ "dependent" "rewrite" orient(b) constr(c) ] -> [ rewriteInConcl b c ] @@ -765,6 +767,6 @@ END;; mode. *) VERNAC COMMAND EXTEND GrabEvars [ "Grab" "Existential" "Variables" ] -> - [ let p = Proof_global.give_me_the_proof () in - Proof.V82.grab_evars p ] + [ Proof_global.with_current_proof (fun _ p -> + Proof.V82.grab_evars p) ] END diff --git a/tactics/leminv.ml b/tactics/leminv.ml index ac4cd5442..63a591777 100644 --- a/tactics/leminv.ml +++ b/tactics/leminv.ml @@ -196,7 +196,8 @@ let inversion_scheme env sigma t sort dep_option inv_op = (str"Computed inversion goal was not closed in initial signature."); *) let pf = Proof.start [invEnv,invGoal] in - Proof.run_tactic env (Proofview.V82.tactic (tclTHEN intro (onLastHypId inv_op))) pf; + let pf = Proof.run_tactic env + (Proofview.V82.tactic (tclTHEN intro (onLastHypId inv_op))) pf in let pfterm = List.hd (Proof.partial_proof pf) in let global_named_context = Global.named_context () in let ownSign = ref begin @@ -227,7 +228,7 @@ let inversion_scheme env sigma t sort dep_option inv_op = let add_inversion_lemma name env sigma t sort dep inv_op = let invProof = inversion_scheme env sigma t sort dep inv_op in let entry = { - const_entry_body = invProof; + const_entry_body = Future.from_val (invProof,Declareops.no_seff); const_entry_secctx = None; const_entry_type = None; const_entry_opaque = false; @@ -241,8 +242,8 @@ let add_inversion_lemma name env sigma t sort dep inv_op = let inversion_lemma_from_goal n na (loc,id) sort dep_option inv_op = let pts = get_pftreestate() in - let { it=gls ; sigma=sigma } = Proof.V82.subgoals pts in - let gl = { it = List.nth gls (n-1) ; sigma=sigma } in + let { it=gls ; sigma=sigma; eff=eff } = Proof.V82.subgoals pts in + let gl = { it = List.nth gls (n-1) ; sigma=sigma; eff=eff } in let t = try pf_get_hyp_typ gl id with Not_found -> Pretype_errors.error_var_not_found_loc loc id in diff --git a/tactics/rewrite.ml4 b/tactics/rewrite.ml4 index c88e67e8b..154cad691 100644 --- a/tactics/rewrite.ml4 +++ b/tactics/rewrite.ml4 @@ -683,7 +683,7 @@ let reset_env env = let fold_match ?(force=false) env sigma c = let (ci, p, c, brs) = destCase c in let cty = Retyping.get_type_of env sigma c in - let dep, pred, exists, sk = + let dep, pred, exists, (sk, eff) = let env', ctx, body = let ctx, pred = decompose_lam_assum p in let env' = Environ.push_rel_context ctx env in @@ -720,7 +720,7 @@ let fold_match ?(force=false) env sigma c = let meths = List.map (fun br -> br) (Array.to_list brs) in applist (mkConst sk, pars @ [pred] @ meths @ args @ [c]) in - sk, (if exists then env else reset_env env), app + sk, (if exists then env else reset_env env), app, eff let unfold_match env sigma sk app = match kind_of_term app with @@ -897,7 +897,7 @@ let subterm all flags (s : strategy) : strategy = else match try Some (fold_match env (goalevars evars) t) with Not_found -> None with | None -> x - | Some (cst, _, t') -> + | Some (cst, _, t',_) -> (* eff XXX *) match aux env avoid t' ty cstr evars with | Some (Some prf) -> Some (Some { prf with @@ -1251,7 +1251,9 @@ let assert_replacing id newt tac = env ~init:[] in let (e, args) = destEvar ev in - Goal.return (mkEvar (e, Array.of_list inst))))) + Goal.return + (mkEvar (e, Array.of_list inst)) + Declareops.no_seff))) in Goal.bind reft Goal.refine) in Proofview.tclTHEN (Proofview.tclSENSITIVE sens) (Proofview.tclFOCUS 2 2 tac) @@ -1295,7 +1297,7 @@ let cl_rewrite_clause_newtac ?abs strat clause = try let res = cl_rewrite_clause_aux ?abs strat env [] sigma ty is_hyp - in return (res, is_hyp) + in return (res, is_hyp) Declareops.no_seff with | TypeClassError (env, (UnsatisfiableConstraints _ as e)) -> raise (RewriteFailure (str"Unable to satisfy the rewriting constraints." @@ -1316,9 +1318,9 @@ let cl_rewrite_clause_new_strat ?abs strat clause = newfail 0 (str"setoid rewrite failed: " ++ s)) let cl_rewrite_clause_newtac' l left2right occs clause = - Proof_global.run_tactic + Proof_global.with_current_proof (fun _ -> Proof.run_tactic (Global.env ()) (Proofview.tclFOCUS 1 1 - (cl_rewrite_clause_new_strat (rewrite_with rewrite_unif_flags l left2right occs) clause)) + (cl_rewrite_clause_new_strat (rewrite_with rewrite_unif_flags l left2right occs) clause))) let cl_rewrite_clause_strat strat clause = tclTHEN (tactic_init_setoid ()) @@ -1722,7 +1724,8 @@ let declare_projection n instance_id r = let ty = Global.type_of_global r in let c = constr_of_global r in let term = proper_projection c ty in - let typ = Typing.type_of (Global.env ()) Evd.empty term in + let env = Global.env() in + let typ = Typing.type_of env Evd.empty term in let ctx, typ = decompose_prod_assum typ in let typ = let n = @@ -1739,12 +1742,12 @@ let declare_projection n instance_id r = | _ -> typ in aux init in - let ctx,ccl = Reductionops.splay_prod_n (Global.env()) Evd.empty (3 * n) typ + let ctx,ccl = Reductionops.splay_prod_n env Evd.empty (3 * n) typ in it_mkProd_or_LetIn ccl ctx in let typ = it_mkProd_or_LetIn typ ctx in let cst = - { const_entry_body = term; + { const_entry_body = Future.from_val (term,Declareops.no_seff); const_entry_secctx = None; const_entry_type = Some typ; const_entry_opaque = false; @@ -2052,26 +2055,30 @@ TACTIC EXTEND implify [ "implify" hyp(n) ] -> [ implify n ] END -let rec fold_matches env sigma c = +let rec fold_matches eff env sigma c = map_constr_with_full_binders Environ.push_rel (fun env c -> match kind_of_term c with | Case _ -> - let cst, env, c' = fold_match ~force:true env sigma c in - fold_matches env sigma c' - | _ -> fold_matches env sigma c) + let cst, env, c',eff' = fold_match ~force:true env sigma c in + eff := Declareops.union_side_effects eff' !eff; + fold_matches eff env sigma c' + | _ -> fold_matches eff env sigma c) env c TACTIC EXTEND fold_match [ "fold_match" constr(c) ] -> [ fun gl -> - let _, _, c' = fold_match ~force:true (pf_env gl) (project gl) c in - change (Some (snd (pattern_of_constr (project gl) c))) c' onConcl gl ] + let _, _, c', eff = fold_match ~force:true (pf_env gl) (project gl) c in + tclTHEN (Tactics.emit_side_effects eff) + (change (Some (snd (pattern_of_constr (project gl) c))) c' onConcl) gl ] END TACTIC EXTEND fold_matches -| [ "fold_matches" constr(c) ] -> [ fun gl -> - let c' = fold_matches (pf_env gl) (project gl) c in - change (Some (snd (pattern_of_constr (project gl) c))) c' onConcl gl ] +| [ "fold_matches" constr(c) ] -> [ fun gl -> + let eff = ref Declareops.no_seff in + let c' = fold_matches eff (pf_env gl) (project gl) c in + tclTHEN (Tactics.emit_side_effects !eff) + (change (Some (snd (pattern_of_constr (project gl) c))) c' onConcl) gl ] END TACTIC EXTEND myapply diff --git a/tactics/tacinterp.ml b/tactics/tacinterp.ml index 1b460b060..d8e03265d 100644 --- a/tactics/tacinterp.ml +++ b/tactics/tacinterp.ml @@ -1046,7 +1046,7 @@ let mk_open_constr_value ist gl c = let mk_hyp_value ist gl c = Value.of_constr (mkVar (interp_hyp ist gl c)) let mk_int_or_var_value ist c = in_gen (topwit wit_int) (interp_int_or_var ist c) -let pack_sigma (sigma,c) = {it=c;sigma=sigma} +let pack_sigma eff (sigma,c) = {it=c;sigma=sigma;eff=eff} let extend_gl_hyps { it=gl ; sigma=sigma } sign = Goal.V82.new_goal_with sigma gl sign @@ -1317,7 +1317,7 @@ and interp_letin ist gl llc u = (* Interprets the Match Context expressions *) and interp_match_goal ist goal lz lr lmr = let (gl,sigma) = Goal.V82.nf_evar (project goal) (sig_it goal) in - let goal = { it = gl ; sigma = sigma } in + let goal = { it = gl ; sigma = sigma; eff = sig_eff goal } in let hyps = pf_hyps goal in let hyps = if lr then List.rev hyps else hyps in let concl = pf_concl goal in @@ -1457,11 +1457,11 @@ and interp_genarg ist gl x = (snd (out_gen (glbwit (wit_open_constr_gen casted)) x))) | ConstrWithBindingsArgType -> in_gen (topwit wit_constr_with_bindings) - (pack_sigma (interp_constr_with_bindings ist (pf_env gl) (project gl) + (pack_sigma (sig_eff gl) (interp_constr_with_bindings ist (pf_env gl) (project gl) (out_gen (glbwit wit_constr_with_bindings) x))) | BindingsArgType -> in_gen (topwit wit_bindings) - (pack_sigma (interp_bindings ist (pf_env gl) (project gl) (out_gen (glbwit wit_bindings) x))) + (pack_sigma (sig_eff gl) (interp_bindings ist (pf_env gl) (project gl) (out_gen (glbwit wit_bindings) x))) | ListArgType ConstrArgType -> let (sigma,v) = interp_genarg_constr_list ist gl x in evdref := sigma; diff --git a/tactics/tactics.ml b/tactics/tactics.ml index 06ad729d4..ee92762cb 100644 --- a/tactics/tactics.ml +++ b/tactics/tactics.ml @@ -3594,6 +3594,9 @@ let abstract_subproof id tac gl = (** ppedrot: seems legit to have abstracted subproofs as local*) let cst = Declare.declare_constant ~internal:Declare.KernelSilent ~local:true id decl in let lem = mkConst cst in + let gl = {gl with eff = + Declareops.cons_side_effects + (Safe_typing.sideff_of_con (Global.safe_env()) cst) gl.eff} in exact_no_check (applist (lem,List.rev (instance_from_named_context sign))) gl @@ -3621,3 +3624,9 @@ let unify ?(state=full_transparent_state) x y gl = let evd = w_unify (pf_env gl) (project gl) Reduction.CONV ~flags x y in tclEVARS evd gl with e when Errors.noncritical e -> tclFAIL 0 (str"Not unifiable") gl + +let emit_side_effects eff gl = + Declareops.iter_side_effects (fun e -> + prerr_endline ("emitting: " ^ Declareops.string_of_side_effect e)) + eff; + {gl with it = [gl.it] ; eff = Declareops.union_side_effects eff gl.eff} diff --git a/tactics/tactics.mli b/tactics/tactics.mli index cdc8eb375..7acd803bd 100644 --- a/tactics/tactics.mli +++ b/tactics/tactics.mli @@ -398,7 +398,10 @@ val general_multi_rewrite : val subst_one : (bool -> Id.t -> Id.t * constr * bool -> tactic) Hook.t + val declare_intro_decomp_eq : ((int -> tactic) -> Coqlib.coq_eq_data * types * (types * constr * constr) -> clausenv -> tactic) -> unit + +val emit_side_effects : Declareops.side_effects -> tactic diff --git a/tools/coqmktop.ml b/tools/coqmktop.ml index 81f0686e0..39a8ede9a 100644 --- a/tools/coqmktop.ml +++ b/tools/coqmktop.ml @@ -30,7 +30,7 @@ let safe_sys_command = cf. dev/dynlink.ml. *) (* 1. Core objects *) -let ocamlobjs = ["str.cma";"unix.cma";"nums.cma";"dynlink.cma"] +let ocamlobjs = ["str.cma";"unix.cma";"nums.cma";"dynlink.cma";"threads.cma"] let camlp4objs = if Coq_config.camlp4 = "camlp5" then ["gramlib.cma"] else ["camlp4lib.cma"] let libobjs = ocamlobjs @ camlp4objs diff --git a/tools/fake_ide.ml b/tools/fake_ide.ml index 14499c228..5bfa377f0 100644 --- a/tools/fake_ide.ml +++ b/tools/fake_ide.ml @@ -11,14 +11,13 @@ exception Comment type coqtop = { - in_chan : in_channel; xml_printer : Xml_printer.t; xml_parser : Xml_parser.t; } let logger level content = () -let eval_call call coqtop = +let base_eval_call call coqtop = prerr_endline (Serialize.pr_call call); let xml_query = Serialize.of_call call in Xml_printer.print coqtop.xml_printer xml_query; @@ -36,18 +35,40 @@ let eval_call call coqtop = in let res = loop () in prerr_endline (Serialize.pr_full_value call res); - match res with Interface.Fail _ -> exit 1 | _ -> () + match res with Interface.Fail _ -> exit 1 | x -> x + +let eval_call c q = ignore(base_eval_call c q) + +let ids = ref [] +let store_id = function + | Interface.Fail _ -> () + | Interface.Good (id, _) -> ids := id :: !ids +let rec erase_ids n = + if n = 0 then + match !ids with + | [] -> Stateid.initial_state_id + | x :: _ -> x + else match !ids with + | id :: rest -> ids := rest; erase_ids (n-1) + | [] -> exit 1 +let eid = ref 0 +let edit () = incr eid; !eid let commands = [ "INTERPRAWSILENT", (fun s -> eval_call (Serialize.interp (0,true,false,s))); "INTERPRAW", (fun s -> eval_call (Serialize.interp (0,true,true,s))); - "INTERPSILENT", (fun s -> eval_call (Serialize.interp (0,false,false,s))); - "INTERP", (fun s -> eval_call (Serialize.interp (0,false,true,s))); - "REWIND", (fun s -> eval_call (Serialize.rewind (int_of_string s))); + "INTERPSILENT", (fun s c -> + store_id (base_eval_call (Serialize.interp (edit(),false,false,s)) c)); + "INTERP", (fun s c -> + store_id (base_eval_call (Serialize.interp (edit(),false,true,s)) c)); + "REWIND", (fun s -> + let i = int_of_string s in + let id = erase_ids i in + eval_call (Serialize.backto id)); "GOALS", (fun _ -> eval_call (Serialize.goals ())); "HINTS", (fun _ -> eval_call (Serialize.hints ())); "GETOPTIONS", (fun _ -> eval_call (Serialize.get_options ())); - "STATUS", (fun _ -> eval_call (Serialize.status ())); + "STATUS", (fun _ -> eval_call (Serialize.status false)); "INLOADPATH", (fun s -> eval_call (Serialize.inloadpath s)); "MKCASES", (fun s -> eval_call (Serialize.mkcases s)); "#", (fun _ -> raise Comment); @@ -91,7 +112,6 @@ let main = let op = Xml_printer.make (Xml_printer.TChannel cout) in let () = Xml_parser.check_eof ip false in { - in_chan = cin; xml_printer = op; xml_parser = ip; } diff --git a/toplevel/auto_ind_decl.ml b/toplevel/auto_ind_decl.ml index 82445a0fd..e63c22bf8 100644 --- a/toplevel/auto_ind_decl.ml +++ b/toplevel/auto_ind_decl.ml @@ -178,23 +178,30 @@ let build_beq_scheme kn = let rec aux c = let (c,a) = Reductionops.whd_betaiota_stack Evd.empty c in match kind_of_term c with - | Rel x -> mkRel (x-nlist+ndx) - | Var x -> mkVar (Id.of_string ("eq_"^(Id.to_string x))) + | Rel x -> mkRel (x-nlist+ndx), Declareops.no_seff + | Var x -> mkVar (id_of_string ("eq_"^(string_of_id x))), Declareops.no_seff | Cast (x,_,_) -> aux (applist (x,a)) | App _ -> assert false - | Ind (kn',i as ind') -> if eq_mind kn kn' then mkRel(eqA-nlist-i+nb_ind-1) - else ( try - let a = Array.of_list a in - let eq = mkConst (find_scheme (!beq_scheme_kind_aux()) (kn',i)) - and eqa = Array.map aux a - in - let args = Array.append - (Array.map (fun x->lift lifti x) a) eqa - in if Array.equal eq_constr args [||] then eq - else mkApp (eq,Array.append - (Array.map (fun x->lift lifti x) a) eqa) - with Not_found -> raise(EqNotFound (ind',ind)) - ) + | Ind (kn',i as ind') -> + if eq_mind kn kn' then mkRel(eqA-nlist-i+nb_ind-1), Declareops.no_seff + else begin + try + let eq, eff = + let c, eff = find_scheme (!beq_scheme_kind_aux()) (kn',i) in + mkConst c, eff in + let eqa, eff = + let eqa, effs = List.split (List.map aux a) in + Array.of_list eqa, + Declareops.union_side_effects + (Declareops.flatten_side_effects (List.rev effs)) + eff in + let args = + Array.append + (Array.of_list (List.map (fun x -> lift lifti x) a)) eqa in + if Int.equal (Array.length args) 0 then eq, eff + else mkApp (eq, args), eff + with Not_found -> raise(EqNotFound (ind',ind)) + end | Sort _ -> raise InductiveWithSort | Prod _ -> raise InductiveWithProduct | Lambda _-> raise (EqUnknown "Lambda") @@ -229,6 +236,7 @@ let build_beq_scheme kn = let constrsi = constrs (3+nparrec) in let n = Array.length constrsi in let ar = Array.make n ff in + let eff = ref Declareops.no_seff in for i=0 to n-1 do let nb_cstr_args = List.length constrsi.(i).cs_args in let ar2 = Array.make n ff in @@ -240,12 +248,13 @@ let build_beq_scheme kn = | _ -> let eqs = Array.make nb_cstr_args tt in for ndx = 0 to nb_cstr_args-1 do let _,_,cc = List.nth constrsi.(i).cs_args ndx in - let eqA = compute_A_equality rel_list + let eqA, eff' = compute_A_equality rel_list nparrec (nparrec+3+2*nb_cstr_args) (nb_cstr_args+ndx+1) cc in + eff := Declareops.union_side_effects eff' !eff; Array.set eqs ndx (mkApp (eqA, [|mkRel (ndx+1+nb_cstr_args);mkRel (ndx+1)|] @@ -271,23 +280,28 @@ let build_beq_scheme kn = done; mkNamedLambda (Id.of_string "X") (mkFullInd ind (nb_ind-1+1)) ( mkNamedLambda (Id.of_string "Y") (mkFullInd ind (nb_ind-1+2)) ( - mkCase (ci, do_predicate rel_list 0,mkVar (Id.of_string "X"),ar))) + mkCase (ci, do_predicate rel_list 0,mkVar (Id.of_string "X"),ar))), + !eff in (* build_beq_scheme *) let names = Array.make nb_ind Anonymous and types = Array.make nb_ind mkSet and cores = Array.make nb_ind mkSet in + let eff = ref Declareops.no_seff in for i=0 to (nb_ind-1) do names.(i) <- Name (Id.of_string (rec_name i)); types.(i) <- mkArrow (mkFullInd (kn,i) 0) (mkArrow (mkFullInd (kn,i) 1) bb); - cores.(i) <- make_one_eq i + let c, eff' = make_one_eq i in + cores.(i) <- c; + eff := Declareops.union_side_effects eff' !eff done; Array.init nb_ind (fun i -> let kelim = Inductive.elim_sorts (mib,mib.mind_packets.(i)) in if not (List.mem InSet kelim) then raise (NonSingletonProp (kn,i)); let fix = mkFix (((Array.make nb_ind 0),i),(names,types,cores)) in - create_input fix) + create_input fix), + !eff let beq_scheme_kind = declare_mutual_scheme_object "_beq" build_beq_scheme @@ -338,8 +352,10 @@ let do_replace_lb lb_scheme_key aavoid narg gls p q = in let type_of_pq = pf_type_of gls p in let u,v = destruct_ind type_of_pq - in let lb_type_of_p = - try mkConst (find_scheme lb_scheme_key u) + in let lb_type_of_p, eff = + try + let c, eff = find_scheme lb_scheme_key u in + mkConst c, eff with Not_found -> (* spiwack: the format of this error message should probably be improved. *) @@ -357,7 +373,9 @@ let do_replace_lb lb_scheme_key aavoid narg gls p q = (Array.map (fun x -> do_arg x 2) v) in let app = if Array.equal eq_constr lb_args [||] then lb_type_of_p else mkApp (lb_type_of_p,lb_args) - in [Equality.replace p q ; apply app ; Auto.default_auto] + in + [Tactics.emit_side_effects eff; + Equality.replace p q ; apply app ; Auto.default_auto] (* used in the bool -> leib side *) let do_replace_bl bl_scheme_key ind gls aavoid narg lft rgt = @@ -396,8 +414,10 @@ let do_replace_bl bl_scheme_key ind gls aavoid narg lft rgt = in if eq_ind u ind then (Equality.replace t1 t2)::(Auto.default_auto)::(aux q1 q2) else ( - let bl_t1 = - try mkConst (find_scheme bl_scheme_key u) + let bl_t1, eff = + try + let c, eff = find_scheme bl_scheme_key u in + mkConst c, eff with Not_found -> (* spiwack: the format of this error message should probably be improved. *) @@ -418,6 +438,7 @@ let do_replace_bl bl_scheme_key ind gls aavoid narg lft rgt = let app = if Array.equal eq_constr bl_args [||] then bl_t1 else mkApp (bl_t1,bl_args) in + (Tactics.emit_side_effects eff):: (Equality.replace_by t1 t2 (tclTHEN (apply app) (Auto.default_auto)))::(aux q1 q2) ) @@ -463,16 +484,17 @@ let eqI ind l = let list_id = list_id l in let eA = Array.of_list((List.map (fun (s,_,_,_) -> mkVar s) list_id)@ (List.map (fun (_,seq,_,_)-> mkVar seq) list_id )) - and e = try mkConst (find_scheme beq_scheme_kind ind) with - Not_found -> error + and e, eff = + try let c, eff = find_scheme beq_scheme_kind ind in mkConst c, eff + with Not_found -> error ("The boolean equality on "^(string_of_mind (fst ind))^" is needed."); - in (if Array.equal eq_constr eA [||] then e else mkApp(e,eA)) + in (if Array.equal eq_constr eA [||] then e else mkApp(e,eA)), eff (**********************************************************************) (* Boolean->Leibniz *) let compute_bl_goal ind lnamesparrec nparrec = - let eqI = eqI ind lnamesparrec in + let eqI, eff = eqI ind lnamesparrec in let list_id = list_id lnamesparrec in let create_input c = let x = Id.of_string "x" and @@ -506,7 +528,7 @@ let compute_bl_goal ind lnamesparrec nparrec = mkArrow (mkApp(eq,[|bb;mkApp(eqI,[|mkVar n;mkVar m|]);tt|])) (mkApp(eq,[|mkFullInd ind (nparrec+3);mkVar n;mkVar m|])) - ))) + ))), eff let compute_bl_tact bl_scheme_key ind lnamesparrec nparrec gsig = let list_id = list_id lnamesparrec in @@ -590,9 +612,10 @@ let make_bl_scheme mind = let nparrec = mib.mind_nparams_rec in let lnonparrec,lnamesparrec = context_chop (nparams-nparrec) mib.mind_params_ctxt in - [|Pfedit.build_by_tactic (Global.env()) - (compute_bl_goal ind lnamesparrec nparrec) - (compute_bl_tact (!bl_scheme_kind_aux()) ind lnamesparrec nparrec)|] + let bl_goal, eff = compute_bl_goal ind lnamesparrec nparrec in + [|Pfedit.build_by_tactic (Global.env()) bl_goal + (compute_bl_tact (!bl_scheme_kind_aux()) ind lnamesparrec nparrec)|], + eff let bl_scheme_kind = declare_mutual_scheme_object "_dec_bl" make_bl_scheme @@ -603,7 +626,7 @@ let _ = bl_scheme_kind_aux := fun () -> bl_scheme_kind let compute_lb_goal ind lnamesparrec nparrec = let list_id = list_id lnamesparrec in - let eqI = eqI ind lnamesparrec in + let eqI, eff = eqI ind lnamesparrec in let create_input c = let x = Id.of_string "x" and y = Id.of_string "y" in @@ -636,7 +659,7 @@ let compute_lb_goal ind lnamesparrec nparrec = mkArrow (mkApp(eq,[|mkFullInd ind (nparrec+2);mkVar n;mkVar m|])) (mkApp(eq,[|bb;mkApp(eqI,[|mkVar n;mkVar m|]);tt|])) - ))) + ))), eff let compute_lb_tact lb_scheme_key ind lnamesparrec nparrec gsig = let list_id = list_id lnamesparrec in @@ -702,9 +725,10 @@ let make_lb_scheme mind = let nparrec = mib.mind_nparams_rec in let lnonparrec,lnamesparrec = context_chop (nparams-nparrec) mib.mind_params_ctxt in - [|Pfedit.build_by_tactic (Global.env()) - (compute_lb_goal ind lnamesparrec nparrec) - (compute_lb_tact (!lb_scheme_kind_aux()) ind lnamesparrec nparrec)|] + let lb_goal, eff = compute_lb_goal ind lnamesparrec nparrec in + [|Pfedit.build_by_tactic (Global.env()) lb_goal + (compute_lb_tact (!lb_scheme_kind_aux()) ind lnamesparrec nparrec)|], + eff let lb_scheme_kind = declare_mutual_scheme_object "_dec_lb" make_lb_scheme @@ -771,7 +795,7 @@ let compute_dec_goal ind lnamesparrec nparrec = let compute_dec_tact ind lnamesparrec nparrec gsig = let list_id = list_id lnamesparrec in - let eqI = eqI ind lnamesparrec in + let eqI, eff = eqI ind lnamesparrec in let avoid = ref [] in let eqtrue x = mkApp(eq,[|bb;x;tt|]) in let eqfalse x = mkApp(eq,[|bb;x;ff|]) in @@ -793,17 +817,22 @@ let compute_dec_tact ind lnamesparrec nparrec gsig = avoid := freshH::(!avoid); let arfresh = Array.of_list fresh_first_intros in let xargs = Array.sub arfresh 0 (2*nparrec) in - let blI = try mkConst (find_scheme bl_scheme_kind ind) with + let blI, eff' = + try let c, eff = find_scheme bl_scheme_kind ind in mkConst c, eff with Not_found -> error ( "Error during the decidability part, boolean to leibniz"^ " equality is required.") in - let lbI = try mkConst (find_scheme lb_scheme_kind ind) with + let lbI, eff'' = + try let c, eff = find_scheme lb_scheme_kind ind in mkConst c, eff with Not_found -> error ( "Error during the decidability part, leibniz to boolean"^ " equality is required.") in tclTHENSEQ [ + Tactics.emit_side_effects + (Declareops.union_side_effects eff'' + (Declareops.union_side_effects eff' eff)); intros_using fresh_first_intros; intros_using [freshn;freshm]; (*we do this so we don't have to prove the same goal twice *) @@ -859,7 +888,7 @@ let make_eq_decidability mind = context_chop (nparams-nparrec) mib.mind_params_ctxt in [|Pfedit.build_by_tactic (Global.env()) (compute_dec_goal ind lnamesparrec nparrec) - (compute_dec_tact ind lnamesparrec nparrec)|] + (compute_dec_tact ind lnamesparrec nparrec)|], Declareops.no_seff let eq_dec_scheme_kind = declare_mutual_scheme_object "_eq_dec" make_eq_decidability diff --git a/toplevel/auto_ind_decl.mli b/toplevel/auto_ind_decl.mli index b05f94c3c..660f6f7e7 100644 --- a/toplevel/auto_ind_decl.mli +++ b/toplevel/auto_ind_decl.mli @@ -30,17 +30,17 @@ exception ParameterWithoutEquality of constant exception NonSingletonProp of inductive val beq_scheme_kind : mutual scheme_kind -val build_beq_scheme : mutual_inductive -> constr array +val build_beq_scheme : mutual_inductive -> constr array * Declareops.side_effects (** {6 Build equivalence between boolean equality and Leibniz equality } *) val lb_scheme_kind : mutual scheme_kind -val make_lb_scheme : mutual_inductive -> constr array +val make_lb_scheme : mutual_inductive -> constr array * Declareops.side_effects val bl_scheme_kind : mutual scheme_kind -val make_bl_scheme : mutual_inductive -> constr array +val make_bl_scheme : mutual_inductive -> constr array * Declareops.side_effects (** {6 Build decidability of equality } *) val eq_dec_scheme_kind : mutual scheme_kind -val make_eq_decidability : mutual_inductive -> constr array +val make_eq_decidability : mutual_inductive -> constr array * Declareops.side_effects diff --git a/toplevel/autoinstance.ml b/toplevel/autoinstance.ml index 049d631bf..e9601c123 100644 --- a/toplevel/autoinstance.ml +++ b/toplevel/autoinstance.ml @@ -187,7 +187,7 @@ let declare_record_instance gr ctx params = let ident = make_instance_ident gr in let def = it_mkLambda_or_LetIn (applistc (constr_of_global gr) params) ctx in let def = deep_refresh_universes def in - let ce = { const_entry_body= def; + let ce = { const_entry_body= Future.from_val (def,Declareops.no_seff); const_entry_secctx = None; const_entry_type=None; const_entry_opaque=false; @@ -204,9 +204,9 @@ let declare_class_instance gr ctx params = let def = deep_refresh_universes def in let typ = deep_refresh_universes typ in let ce = Entries.DefinitionEntry - { const_entry_type = Some typ; + { const_entry_type= Some typ; + const_entry_body= Future.from_val (def,Declareops.no_seff); const_entry_secctx = None; - const_entry_body = def; const_entry_opaque = false; const_entry_inline_code = false } in try diff --git a/toplevel/backtrack.ml b/toplevel/backtrack.ml deleted file mode 100644 index e259da7af..000000000 --- a/toplevel/backtrack.ml +++ /dev/null @@ -1,234 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* l:=i::!l) history; - !l - -(** Basic manipulation of the command history stack *) - -exception Invalid - -let pop () = ignore (Stack.pop history) - -let npop n = - (* Since our history stack always contains an initial entry, - it's invalid to try to completely empty it *) - if n < 0 || n >= Stack.length history then raise Invalid - else for _i = 1 to n do pop () done - -let top () = - try Stack.top history with Stack.Empty -> raise Invalid - -(** Search the history stack for a suitable location. We perform first - a non-destructive search: in case of search failure, the stack is - unchanged. *) - -exception Found of info - -let search test = - try - Stack.iter (fun i -> if test i then raise (Found i)) history; - raise Invalid - with Found i -> - while i != Stack.top history do pop () done - -(** An auxiliary function to retrieve the number of remaining subgoals *) - -let get_ngoals () = - try - let prf = Proof_global.give_me_the_proof () in - List.length (Evd.sig_it (Proof.V82.background_subgoals prf)) - with Proof_global.NoCurrentProof -> 0 - -(** Register the end of a command and store the current state *) - -let mark_command ast = - Lib.add_frozen_state(); - Lib.mark_end_of_command(); - Stack.push - { label = Lib.current_command_label (); - nproofs = List.length (Pfedit.get_all_proof_names ()); - prfname = - (try Some (Pfedit.get_current_proof_name ()) - with Proof_global.NoCurrentProof -> None); - prfdepth = max 0 (Pfedit.current_proof_depth ()); - reachable = true; - ngoals = get_ngoals (); - cmd = ast } - history - -(** Backtrack by aborting [naborts] proofs, then setting proof-depth back to - [pnum] and finally going to state number [snum]. *) - -let raw_backtrack snum pnum naborts = - for _i = 1 to naborts do Pfedit.delete_current_proof () done; - Pfedit.undo_todepth pnum; - Lib.reset_label snum - -(** Re-sync the state of the system (label, proofs) with the top - of the history stack. We may end on some earlier state to avoid - re-opening proofs. This function will return the final label - and the number of extra backtracking steps performed. *) - -let sync nb_opened_proofs = - (* Backtrack by enough additional steps to avoid re-opening proofs. - Typically, when a Qed has been crossed, we backtrack to the proof start. - NB: We cannot reach the empty stack, since the first entry in the - stack has no opened proofs and is tagged as reachable. - *) - let extra = ref 0 in - while not (top()).reachable do incr extra; pop () done; - let target = top () - in - (* Now the opened proofs at target is a subset of the opened proofs before - the backtrack, we simply abort the extra proofs (if any). - NB: It is critical here that proofs are nested in a regular way - (i.e. no more Resume or Suspend commands as earlier). This way, we can - simply count the extra proofs to abort instead of taking care of their - names. - *) - let naborts = nb_opened_proofs - target.nproofs in - (* We are now ready to do a low-level backtrack *) - raw_backtrack target.label target.prfdepth naborts; - (target.label, !extra) - -(** Backtracking by a certain number of (non-state-preserving) commands. - This is used by Coqide. - It may actually undo more commands than asked : for instance instead - of jumping back in the middle of a finished proof, we jump back before - this proof. The number of extra backtracked command is returned at - the end. *) - -let back count = - if Int.equal count 0 then 0 - else - let nb_opened_proofs = List.length (Pfedit.get_all_proof_names ()) in - npop count; - snd (sync nb_opened_proofs) - -(** Backtracking to a certain state number, and reset proofs accordingly. - We may end on some earlier state if needed to avoid re-opening proofs. - Return the final state number. *) - -let backto snum = - if Int.equal snum (Lib.current_command_label ()) then snum - else - let nb_opened_proofs = List.length (Pfedit.get_all_proof_names ()) in - search (fun i -> Int.equal i.label snum); - fst (sync nb_opened_proofs) - -(** Old [Backtrack] code with corresponding update of the history stack. - [Backtrack] is now deprecated (in favor of [BackTo]) but is kept for - compatibility with ProofGeneral. It's completely up to ProofGeneral - to decide where to go and how to adapt proofs. Note that the choices - of ProofGeneral are currently not always perfect (for instance when - backtracking an Undo). *) - -let backtrack snum pnum naborts = - raw_backtrack snum pnum naborts; - search (fun i -> Int.equal i.label snum) - -(** [reset_initial] resets the system and clears the command history - stack, only pushing back the initial entry. It should be equivalent - to [backto n0] where [n0] is the first label stored in the history. - Note that there might be other labels before [n0] in the libstack, - created during evaluation of .coqrc or initial Load's. *) - -let reset_initial () = - let n = Stack.length history in - npop (n-1); - Pfedit.delete_all_proofs (); - Lib.reset_label (top ()).label - -(** Reset to the last known state just before defining [id] *) - -let reset_name id = - let lbl = - try Lib.label_before_name id with Not_found -> raise Invalid - in - ignore (backto lbl) - -(** When a proof is ended (via either Qed/Admitted/Restart/Abort), - old proof steps should be marked differently to avoid jumping back - to them: - - either this proof isn't there anymore in the proof engine - - either it's there but it's a more recent attempt after a Restart, - so we shouldn't mix the two. - We also mark as unreachable the proof steps cancelled via a Undo. *) - -let mark_unreachable ?(after=0) prf_lst = - let fix i = match i.prfname with - | None -> raise Not_found (* stop hacking the history outside of proofs *) - | Some p -> - if List.mem p prf_lst && i.prfdepth > after - then i.reachable <- false - in - try Stack.iter fix history with Not_found -> () - -(** Parse the history stack for printing the script of a proof *) - -let get_script prf = - let script = ref [] in - let select i = match i.prfname with - | None -> raise Not_found - | Some p when Id.equal p prf && i.reachable -> script := i :: !script - | _ -> () - in - (try Stack.iter select history with Not_found -> ()); - (* Get rid of intermediate commands which don't grow the proof depth *) - let rec filter n = function - | [] -> [] - | {prfdepth=d; cmd=c; ngoals=ng}::l when n < d -> (c,ng) :: filter d l - | {prfdepth=d}::l -> filter d l - in - (* initial proof depth (after entering the lemma statement) is 1 *) - filter 1 !script diff --git a/toplevel/backtrack.mli b/toplevel/backtrack.mli deleted file mode 100644 index d4ac7cc61..000000000 --- a/toplevel/backtrack.mli +++ /dev/null @@ -1,101 +0,0 @@ -(************************************************************************) -(* v * The Coq Proof Assistant / The Coq Development Team *) -(* unit - -(** Is this history stack active (i.e. nonempty) ? - The stack is currently inactive when compiling files (coqc). *) - -val is_active : unit -> bool - -(** The [Invalid] exception is raised when one of the following function - tries to empty the history stack, or reach an unknown states, etc. - The stack is preserved in these cases. *) - -exception Invalid - -(** Nota Bene: it is critical for the following functions that proofs - are nested in a regular way (i.e. no more Resume or Suspend commands - as earlier). *) - -(** Backtracking by a certain number of (non-state-preserving) commands. - This is used by Coqide. - It may actually undo more commands than asked : for instance instead - of jumping back in the middle of a finished proof, we jump back before - this proof. The number of extra backtracked command is returned at - the end. *) - -val back : int -> int - -(** Backtracking to a certain state number, and reset proofs accordingly. - We may end on some earlier state if needed to avoid re-opening proofs. - Return the state number on which we finally end. *) - -val backto : int -> int - -(** Old [Backtrack] code with corresponding update of the history stack. - [Backtrack] is now deprecated (in favor of [BackTo]) but is kept for - compatibility with ProofGeneral. It's completely up to ProofGeneral - to decide where to go and how to adapt proofs. Note that the choices - of ProofGeneral are currently not always perfect (for instance when - backtracking an Undo). *) - -val backtrack : int -> int -> int -> unit - -(** [reset_initial] resets the system and clears the command history - stack, only pushing back the initial entry. It should be equivalent - to [backto n0] where [n0] is the first label stored in the history. - Note that there might be other labels before [n0] in the libstack, - created during evaluation of .coqrc or initial Load's. *) - -val reset_initial : unit -> unit - -(** Reset to the last known state just before defining [id] *) - -val reset_name : Names.Id.t Loc.located -> unit - -(** When a proof is ended (via either Qed/Admitted/Restart/Abort), - old proof steps should be marked differently to avoid jumping back - to them: - - either this proof isn't there anymore in the proof engine - - either a proof with the same name is there, but it's a more recent - attempt after a Restart/Abort, we shouldn't mix the two. - We also mark as unreachable the proof steps cancelled via a Undo. -*) - -val mark_unreachable : ?after:int -> Names.Id.t list -> unit - -(** Parse the history stack for printing the script of a proof *) - -val get_script : Names.Id.t -> (Vernacexpr.vernac_expr * int) list - - -(** For debug purpose, a dump of the history *) - -type info = { - label : int; - nproofs : int; - prfname : Names.Id.t option; - prfdepth : int; - ngoals : int; - cmd : Vernacexpr.vernac_expr; - mutable reachable : bool; -} - -val dump_history : unit -> info list diff --git a/toplevel/class.ml b/toplevel/class.ml index 14473e46f..f9ce75bba 100644 --- a/toplevel/class.ml +++ b/toplevel/class.ml @@ -217,7 +217,8 @@ let build_id_coercion idf_opt source = in let constr_entry = (* Cast is necessary to express [val_f] is identity *) DefinitionEntry - { const_entry_body = mkCast (val_f, DEFAULTcast, typ_f); + { const_entry_body = Future.from_val + (mkCast (val_f, DEFAULTcast, typ_f),Declareops.no_seff); const_entry_secctx = None; const_entry_type = Some typ_f; const_entry_opaque = false; diff --git a/toplevel/classes.ml b/toplevel/classes.ml index e4064049e..a217abbba 100644 --- a/toplevel/classes.ml +++ b/toplevel/classes.ml @@ -103,7 +103,7 @@ let instance_hook k pri global imps ?hook cst = let declare_instance_constant k pri global imps ?hook id term termtype = let kind = IsDefinition Instance in let entry = { - const_entry_body = term; + const_entry_body = Future.from_val term; const_entry_secctx = None; const_entry_type = Some termtype; const_entry_opaque = false; @@ -273,7 +273,8 @@ let new_instance ?(abstract=false) ?(global=false) ctx (instid, bk, cl) props let evm = Evarutil.nf_evar_map_undefined !evars in let evm = undefined_evars evm in if Evd.is_empty evm && not (Option.is_empty term) then - declare_instance_constant k pri global imps ?hook id (Option.get term) termtype + declare_instance_constant k pri global imps ?hook id + (Option.get term,Declareops.no_seff) termtype else begin let kind = Decl_kinds.Global, Decl_kinds.DefinitionBody Decl_kinds.Instance in if Flags.is_program_mode () then diff --git a/toplevel/classes.mli b/toplevel/classes.mli index 1d9ae6ec4..5a0dc97c2 100644 --- a/toplevel/classes.mli +++ b/toplevel/classes.mli @@ -42,7 +42,7 @@ val declare_instance_constant : Impargs.manual_explicitation list -> (** implicits *) ?hook:(Globnames.global_reference -> unit) -> Id.t -> (** name *) - Term.constr -> (** body *) + Entries.proof_output -> (** body *) Term.types -> (** type *) Names.Id.t diff --git a/toplevel/command.ml b/toplevel/command.ml index 369abd3f8..d730d8ef1 100644 --- a/toplevel/command.ml +++ b/toplevel/command.ml @@ -67,9 +67,11 @@ let rec complete_conclusion a cs = function let red_constant_entry n ce = function | None -> ce | Some red -> - let body = ce.const_entry_body in - { ce with const_entry_body = - under_binders (Global.env()) (fst (reduction_of_red_expr red)) n body } + let proof_out = ce.const_entry_body in + { ce with const_entry_body = Future.chain ~id:"red_constant_entry" + proof_out (fun (body,eff) -> + under_binders (Global.env()) + (fst (reduction_of_red_expr red)) n body,eff) } let interp_definition bl red_option c ctypopt = let env = Global.env() in @@ -82,7 +84,7 @@ let interp_definition bl red_option c ctypopt = let c, imps2 = interp_constr_evars_impls ~impls evdref env_bl c in let body = nf_evar !evdref (it_mkLambda_or_LetIn c ctx) in imps1@(Impargs.lift_implicits nb_args imps2), - { const_entry_body = body; + { const_entry_body = Future.from_val (body,Declareops.no_seff); const_entry_secctx = None; const_entry_type = None; const_entry_opaque = false; @@ -100,7 +102,7 @@ let interp_definition bl red_option c ctypopt = then msg_warning (strbrk "Implicit arguments declaration relies on type." ++ spc () ++ strbrk "The term declares more implicits than the type here."); imps1@(Impargs.lift_implicits nb_args impsty), - { const_entry_body = body; + { const_entry_body = Future.from_val(body,Declareops.no_seff); const_entry_secctx = None; const_entry_type = Some typ; const_entry_opaque = false; @@ -135,7 +137,7 @@ let declare_definition_hook = ref ignore let set_declare_definition_hook = (:=) declare_definition_hook let get_declare_definition_hook () = !declare_definition_hook -let declare_definition ident (local, k) ce imps hook = +let declare_definition ident (local,k) ce imps hook = let () = !declare_definition_hook ce in let r = match local with | Discharge when Lib.sections_are_opened () -> @@ -158,7 +160,8 @@ let do_definition ident k bl red_option c ctypopt hook = let (ce, evd, imps as def) = interp_definition bl red_option c ctypopt in if Flags.is_program_mode () then let env = Global.env () in - let c = ce.const_entry_body in + let c,sideff = Future.force ce.const_entry_body in + assert(sideff = Declareops.no_seff); let typ = match ce.const_entry_type with | Some t -> t | None -> Retyping.get_type_of env evd c @@ -173,7 +176,7 @@ let do_definition ident k bl red_option c ctypopt hook = (* 2| Variable/Hypothesis/Parameter/Axiom declarations *) -let declare_assumption is_coe (local,kind) c imps impl nl (_,ident) = match local with +let declare_assumption is_coe (local, kind) c imps impl nl (_,ident) = match local with | Discharge when Lib.sections_are_opened () -> let decl = (Lib.cwd(), SectionLocalAssum (c,impl), IsAssumption kind) in let _ = declare_variable ident decl in @@ -521,7 +524,7 @@ let build_fix_type (_,ctx) ccl = it_mkProd_or_LetIn ccl ctx let declare_fix kind f def t imps = let ce = { - const_entry_body = def; + const_entry_body = Future.from_val def; const_entry_secctx = None; const_entry_type = Some t; const_entry_opaque = false; @@ -712,7 +715,7 @@ let build_wellfounded (recname,n,bl,arityc,body) r measure notation = let body = it_mkLambda_or_LetIn (mkApp (constr_of_global gr, [|make|])) binders_rel in let ty = it_mkProd_or_LetIn top_arity binders_rel in let ce = - { const_entry_body = Evarutil.nf_evar !evdref body; + { const_entry_body = Future.from_val (Evarutil.nf_evar !evdref body,Declareops.no_seff); const_entry_secctx = None; const_entry_type = Some ty; const_entry_opaque = false; @@ -824,10 +827,12 @@ let declare_fixpoint local ((fixnames,fixdefs,fixtypes),fiximps) indexes ntns = (* We shortcut the proof process *) let fixdefs = List.map Option.get fixdefs in let fixdecls = prepare_recursive_declaration fixnames fixtypes fixdefs in - let indexes = search_guard Loc.ghost (Global.env()) indexes fixdecls in + let env = Global.env() in + let indexes = search_guard Loc.ghost env indexes fixdecls in let fiximps = List.map (fun (n,r,p) -> r) fiximps in let fixdecls = List.map_i (fun i _ -> mkFix ((indexes,i),fixdecls)) 0 fixnames in + let fixdecls = List.map (fun c -> c, Declareops.no_seff) fixdecls in ignore (List.map4 (declare_fix (local, Fixpoint)) fixnames fixdecls fixtypes fiximps); (* Declare the recursive definitions *) fixpoint_message (Some indexes) fixnames; @@ -850,6 +855,7 @@ let declare_cofixpoint local ((fixnames,fixdefs,fixtypes),fiximps) ntns = 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 fixdecls = List.map (fun c-> c,Declareops.no_seff) fixdecls in let fiximps = List.map (fun (len,imps,idx) -> imps) fiximps in ignore (List.map4 (declare_fix (local, CoFixpoint)) fixnames fixdecls fixtypes fiximps); (* Declare the recursive definitions *) diff --git a/toplevel/command.mli b/toplevel/command.mli index 71eec3295..76a3c5802 100644 --- a/toplevel/command.mli +++ b/toplevel/command.mli @@ -150,4 +150,4 @@ val do_cofixpoint : val check_mutuality : Environ.env -> bool -> (Id.t * types) list -> unit val declare_fix : definition_kind -> Id.t -> - constr -> types -> Impargs.manual_implicits -> global_reference + Entries.proof_output -> types -> Impargs.manual_implicits -> global_reference diff --git a/toplevel/coqtop.ml b/toplevel/coqtop.ml index cb587a6ef..1dba38001 100644 --- a/toplevel/coqtop.ml +++ b/toplevel/coqtop.ml @@ -282,7 +282,7 @@ let parse_args arglist = | "-debug" :: rem -> set_debug (); parse rem - | "-time" :: rem -> Vernac.time := true; parse rem + | "-time" :: rem -> Flags.time := true; parse rem | "-compat" :: v :: rem -> Flags.compat_version := get_compat_version v; parse rem @@ -390,6 +390,7 @@ let init arglist = load_vernac_obj (); require (); load_rcfile(); + Stm.init (); load_vernacular (); compile_files (); outputstate () @@ -407,9 +408,7 @@ let init arglist = Pp.ppnl (with_option raw_print Prettyp.print_full_pure_context ()); Profile.print_profile (); exit 0 - end; - (* We initialize the command history stack with a first entry *) - Backtrack.mark_command Vernacexpr.VernacNop + end let init_toplevel = init diff --git a/toplevel/ide_slave.ml b/toplevel/ide_slave.ml index 709062f1f..821830203 100644 --- a/toplevel/ide_slave.ml +++ b/toplevel/ide_slave.ml @@ -104,10 +104,8 @@ let coqide_cmd_checks (loc,ast) = if is_debug ast then user_error "Debug mode not available within CoqIDE"; if is_known_option ast then - user_error "Use CoqIDE display menu instead"; - if Vernac.is_navigation_vernac ast then - user_error "Use CoqIDE navigation instead"; - if is_undo ast then + msg_warning (strbrk"This will not work. Use CoqIDE display menu instead"); + if Vernac.is_navigation_vernac ast || is_undo ast then msg_warning (strbrk "Rather use CoqIDE navigation instead"); if is_query ast then msg_warning (strbrk "Query commands should not be inserted in scripts") @@ -120,10 +118,14 @@ let interp (id,raw,verbosely,s) = let loc_ast = Vernac.parse_sentence (pa,None) in if not raw then coqide_cmd_checks loc_ast; Flags.make_silent (not verbosely); - Vernac.eval_expr ~preserving:raw loc_ast; + if id = 0 then Vernac.eval_expr (loc, VernacStm (Command ast)) + else Vernac.eval_expr loc_ast; Flags.make_silent true; - Pp.feedback Interface.Processed; - read_stdout () + Stm.get_current_state (), read_stdout () + +let backto id = + Vernac.eval_expr (Loc.ghost, + VernacStm (Command (VernacBackTo (Stateid.int_of_state_id id)))) (** Goal display *) @@ -252,7 +254,7 @@ let status () = Interface.status_proofname = proof; Interface.status_allproofs = allproofs; Interface.status_statenum = Lib.current_command_label (); - Interface.status_proofnum = Pfedit.current_proof_depth (); + Interface.status_proofnum = Stm.current_proof_depth (); } let search flags = Search.interface_search flags @@ -277,6 +279,17 @@ let about () = { Interface.compile_date = Coq_config.compile_date; } +let handle_exn e = + let dummy = Stateid.dummy_state_id in + let loc_of e = match Loc.get_loc e with + | Some loc when not (Loc.is_ghost loc) -> Some (Loc.unloc loc) + | _ -> None in + let mk_msg e = read_stdout ()^"\n"^string_of_ppcmds (Errors.print e) in + match e with + | Errors.Drop -> dummy, None, "Drop is not allowed by coqide!" + | Errors.Quit -> dummy, None, "Quit is not allowed by coqide!" + | e -> dummy, loc_of e, mk_msg e + (** When receiving the Quit call, we don't directly do an [exit 0], but rather set this reference, in order to send a final answer before exiting. *) @@ -307,7 +320,7 @@ let eval_call c = in let handler = { Interface.interp = interruptible interp; - Interface.rewind = interruptible Backtrack.back; + Interface.rewind = interruptible (fun _ -> 0); Interface.goals = interruptible goals; Interface.evars = interruptible evars; Interface.hints = interruptible hints; diff --git a/toplevel/ind_tables.ml b/toplevel/ind_tables.ml index f55a12f9d..0a5f7d25f 100644 --- a/toplevel/ind_tables.ml +++ b/toplevel/ind_tables.ml @@ -27,8 +27,8 @@ open Decl_kinds (**********************************************************************) (* Registering schemes in the environment *) -type mutual_scheme_object_function = mutual_inductive -> constr array -type individual_scheme_object_function = inductive -> constr +type mutual_scheme_object_function = mutual_inductive -> constr array * Declareops.side_effects +type individual_scheme_object_function = inductive -> constr * Declareops.side_effects type 'a scheme_kind = string @@ -67,8 +67,8 @@ type individual type mutual type scheme_object_function = - | MutualSchemeFunction of (mutual_inductive -> constr array) - | IndividualSchemeFunction of (inductive -> constr) + | MutualSchemeFunction of (mutual_inductive -> constr array * Declareops.side_effects) + | IndividualSchemeFunction of (inductive -> constr * Declareops.side_effects) let scheme_object_table = (Hashtbl.create 17 : (string, string * scheme_object_function) Hashtbl.t) @@ -113,7 +113,7 @@ let define internal id c = let fd = declare_constant ~internal in let id = compute_name internal id in let entry = { - const_entry_body = c; + const_entry_body = Future.from_val (c,Declareops.no_seff); const_entry_secctx = None; const_entry_type = None; const_entry_opaque = false; @@ -127,14 +127,14 @@ let define internal id c = kn let define_individual_scheme_base kind suff f internal idopt (mind,i as ind) = - let c = f ind in + let c, eff = f ind in let mib = Global.lookup_mind mind in let id = match idopt with | Some id -> id | None -> add_suffix mib.mind_packets.(i).mind_typename suff in let const = define internal id c in declare_scheme kind [|ind,const|]; - const + const, Declareops.cons_side_effects (Safe_typing.sideff_of_con (Global.safe_env()) const) eff let define_individual_scheme kind internal names (mind,i as ind) = match Hashtbl.find scheme_object_table kind with @@ -143,14 +143,19 @@ let define_individual_scheme kind internal names (mind,i as ind) = define_individual_scheme_base kind s f internal names ind let define_mutual_scheme_base kind suff f internal names mind = - let cl = f mind in + let cl, eff = f mind in let mib = Global.lookup_mind mind in let ids = Array.init (Array.length mib.mind_packets) (fun i -> try List.assoc i names with Not_found -> add_suffix mib.mind_packets.(i).mind_typename suff) in let consts = Array.map2 (define internal) ids cl in declare_scheme kind (Array.mapi (fun i cst -> ((mind,i),cst)) consts); - consts + consts, + Declareops.union_side_effects eff + (Declareops.side_effects_of_list + (List.map + (fun x -> Safe_typing.sideff_of_con (Global.safe_env()) x) + (Array.to_list consts))) let define_mutual_scheme kind internal names mind = match Hashtbl.find scheme_object_table kind with @@ -159,13 +164,16 @@ let define_mutual_scheme kind internal names mind = define_mutual_scheme_base kind s f internal names mind let find_scheme kind (mind,i as ind) = - try String.Map.find kind (Indmap.find ind !scheme_map) + try + let s = String.Map.find kind (Indmap.find ind !scheme_map) in + s, Declareops.no_seff with Not_found -> match Hashtbl.find scheme_object_table kind with | s,IndividualSchemeFunction f -> define_individual_scheme_base kind s f KernelSilent None ind | s,MutualSchemeFunction f -> - (define_mutual_scheme_base kind s f KernelSilent [] mind).(i) + let ca, eff = define_mutual_scheme_base kind s f KernelSilent [] mind in + ca.(i), eff let check_scheme kind ind = try let _ = String.Map.find kind (Indmap.find ind !scheme_map) in true diff --git a/toplevel/ind_tables.mli b/toplevel/ind_tables.mli index 9dd25c63e..c5e82e2e2 100644 --- a/toplevel/ind_tables.mli +++ b/toplevel/ind_tables.mli @@ -22,8 +22,10 @@ type mutual type individual type 'a scheme_kind -type mutual_scheme_object_function = mutual_inductive -> constr array -type individual_scheme_object_function = inductive -> constr +type mutual_scheme_object_function = + mutual_inductive -> constr array * Declareops.side_effects +type individual_scheme_object_function = + inductive -> constr * Declareops.side_effects (** Main functions to register a scheme builder *) @@ -31,7 +33,8 @@ val declare_mutual_scheme_object : string -> ?aux:string -> mutual_scheme_object_function -> mutual scheme_kind val declare_individual_scheme_object : string -> ?aux:string -> - individual_scheme_object_function -> individual scheme_kind + individual_scheme_object_function -> + individual scheme_kind (* val declare_scheme : 'a scheme_kind -> (inductive * constant) array -> unit @@ -41,12 +44,12 @@ val declare_scheme : 'a scheme_kind -> (inductive * constant) array -> unit val define_individual_scheme : individual scheme_kind -> Declare.internal_flag (** internal *) -> - Id.t option -> inductive -> constant + Id.t option -> inductive -> constant * Declareops.side_effects val define_mutual_scheme : mutual scheme_kind -> Declare.internal_flag (** internal *) -> - (int * Id.t) list -> mutual_inductive -> constant array + (int * Id.t) list -> mutual_inductive -> constant array * Declareops.side_effects (** Main function to retrieve a scheme in the cache or to generate it *) -val find_scheme : 'a scheme_kind -> inductive -> constant +val find_scheme : 'a scheme_kind -> inductive -> constant * Declareops.side_effects val check_scheme : 'a scheme_kind -> inductive -> bool diff --git a/toplevel/indschemes.ml b/toplevel/indschemes.ml index 8551c2038..bf07156f7 100644 --- a/toplevel/indschemes.ml +++ b/toplevel/indschemes.ml @@ -355,7 +355,8 @@ let do_mutual_induction_scheme lnamedepindsort = let declare decl fi lrecref = let decltype = Retyping.get_type_of env0 Evd.empty decl in let decltype = refresh_universes decltype in - let cst = define fi UserVerbose decl (Some decltype) in + let proof_output = Future.from_val (decl,Declareops.no_seff) in + let cst = define fi UserVerbose proof_output (Some decltype) in ConstRef cst :: lrecref in let _ = List.fold_right2 declare listdecl lrecnames [] in @@ -449,8 +450,10 @@ let do_combined_scheme name schemes = with Not_found -> error ((string_of_qualid (snd qualid))^" is not declared.")) schemes in - let body,typ = build_combined_scheme (Global.env ()) csts in - ignore (define (snd name) UserVerbose body (Some typ)); + let env = Global.env () in + let body,typ = build_combined_scheme env csts in + let proof_output = Future.from_val (body,Declareops.no_seff) in + ignore (define (snd name) UserVerbose proof_output (Some typ)); fixpoint_message None [snd name] (**********************************************************************) diff --git a/toplevel/lemmas.ml b/toplevel/lemmas.ml index 695426d56..6876483a0 100644 --- a/toplevel/lemmas.ml +++ b/toplevel/lemmas.ml @@ -48,17 +48,22 @@ let adjust_guardness_conditions const = function | [] -> const (* Not a recursive statement *) | possible_indexes -> (* Try all combinations... not optimal *) - match kind_of_term const.const_entry_body with - | Fix ((nv,0),(_,_,fixdefs as fixdecls)) -> + (* XXX bug ignore(Future.join const.const_entry_body); *) + { const with const_entry_body = + Future.chain ~id:"adjust_guardness_conditions" const.const_entry_body + (fun (body, eff) -> + match kind_of_term body with + | Fix ((nv,0),(_,_,fixdefs as fixdecls)) -> (* let possible_indexes = List.map2 (fun i c -> match i with Some i -> i | None -> List.interval 0 (List.length ((lam_assum c)))) lemma_guard (Array.to_list fixdefs) in *) - let indexes = - search_guard Loc.ghost (Global.env()) possible_indexes fixdecls in - { const with const_entry_body = mkFix ((indexes,0),fixdecls) } - | c -> const + let indexes = + search_guard Loc.ghost (Global.env()) + possible_indexes fixdecls in + mkFix ((indexes,0),fixdecls), eff + | _ -> body, eff) } let find_mutually_recursive_statements thms = let n = List.length thms in @@ -159,7 +164,7 @@ let look_for_possibly_mutual_statements = function (* Saving a goal *) -let save id const do_guard (locality,kind) hook = +let save ?proof id const do_guard (locality,kind) hook = let const = adjust_guardness_conditions const do_guard in let k = Kindops.logical_kind_of_goal_kind kind in let l,r = match locality with @@ -175,7 +180,8 @@ let save id const do_guard (locality,kind) hook = let kn = declare_constant id ~local (DefinitionEntry const, k) in Autoinstance.search_declaration (ConstRef kn); (locality, ConstRef kn) in - Pfedit.delete_current_proof (); + (* if the proof is given explicitly, nothing has to be deleted *) + if proof = None then Pfedit.delete_current_proof (); definition_message id; hook l r @@ -220,7 +226,8 @@ let save_remaining_recthms (locality,kind) body opaq i (id,(t_i,(_,imps))) = | _ -> anomaly (Pp.str "Not a proof by induction") in match locality with | Discharge -> - let const = { const_entry_body = body_i; + let const = { const_entry_body = + Future.from_val (body_i,Declareops.no_seff); const_entry_secctx = None; const_entry_type = Some t_i; const_entry_opaque = opaq; @@ -235,7 +242,8 @@ let save_remaining_recthms (locality,kind) body opaq i (id,(t_i,(_,imps))) = | Global -> false | Discharge -> assert false in - let const = { const_entry_body = body_i; + let const = { const_entry_body = + Future.from_val (body_i,Declareops.no_seff); const_entry_secctx = None; const_entry_type = Some t_i; const_entry_opaque = opaq; @@ -247,37 +255,32 @@ let save_remaining_recthms (locality,kind) body opaq i (id,(t_i,(_,imps))) = let save_hook = ref ignore let set_save_hook f = save_hook := f -let get_proof opacity = - let id,(const,do_guard,persistence,hook) = Pfedit.cook_proof !save_hook in +let get_proof ?proof opacity = + let id,(const,do_guard,persistence,hook) = + match proof with + | None -> Pfedit.cook_proof !save_hook + | Some p -> Pfedit.cook_this_proof !save_hook p in id,{const with const_entry_opaque = opacity},do_guard,persistence,hook -let save_named opacity = - let p = Proof_global.give_me_the_proof () in - Proof.transaction p begin fun () -> - let id,const,do_guard,persistence,hook = get_proof opacity in - save id const do_guard persistence hook - end +let save_named ?proof opacity = + let id,const,do_guard,persistence,hook = get_proof ?proof opacity in + save ?proof id const 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 error "This command can only be used for unnamed theorem." -let save_anonymous opacity save_ident = - let p = Proof_global.give_me_the_proof () in - Proof.transaction p begin fun () -> - let id,const,do_guard,persistence,hook = get_proof opacity in - check_anonymity id save_ident; - save save_ident const do_guard persistence hook - end -let save_anonymous_with_strength kind opacity save_ident = - let p = Proof_global.give_me_the_proof () in - Proof.transaction p begin fun () -> - let id,const,do_guard,_,hook = get_proof opacity in - check_anonymity id save_ident; - (* we consider that non opaque behaves as local for discharge *) - save save_ident const do_guard (Global, Proof kind) hook - end +let save_anonymous ?proof opacity save_ident = + let id,const,do_guard,persistence,hook = get_proof ?proof opacity in + check_anonymity id save_ident; + save ?proof save_ident const do_guard persistence hook + +let save_anonymous_with_strength ?proof kind opacity save_ident = + let id,const,do_guard,_,hook = get_proof ?proof opacity in + check_anonymity id save_ident; + (* we consider that non opaque behaves as local for discharge *) + save ?proof save_ident const do_guard (Global, Proof kind) hook (* Starting a goal *) diff --git a/toplevel/lemmas.mli b/toplevel/lemmas.mli index d6bc90bc3..e8998d608 100644 --- a/toplevel/lemmas.mli +++ b/toplevel/lemmas.mli @@ -35,22 +35,24 @@ val start_proof_with_initialization : val set_save_hook : (Proof.proof -> unit) -> unit (** {6 ... } *) -(** [save_named b] saves the current completed proof under the name it -was started; boolean [b] tells if the theorem is declared opaque; it -fails if the proof is not completed *) +(** [save_named b] saves the current completed (or the provided) proof + under the name it was started; boolean [b] tells if the theorem is + declared opaque; it fails if the proof is not completed *) -val save_named : bool -> unit +val save_named : ?proof:Proof_global.closed_proof -> bool -> unit (** [save_anonymous b name] behaves as [save_named] but declares the theorem under the name [name] and respects the strength of the declaration *) -val save_anonymous : bool -> Id.t -> unit +val save_anonymous : + ?proof:Proof_global.closed_proof -> bool -> Id.t -> unit (** [save_anonymous_with_strength s b name] behaves as [save_anonymous] but declares the theorem under the name [name] and gives it the strength [strength] *) -val save_anonymous_with_strength : theorem_kind -> bool -> Id.t -> unit +val save_anonymous_with_strength : + ?proof:Proof_global.closed_proof -> theorem_kind -> bool -> Id.t -> unit (** [admit ()] aborts the current goal and save it as an assmumption *) diff --git a/toplevel/obligations.ml b/toplevel/obligations.ml index 44ff40f24..9e7ddd5a6 100644 --- a/toplevel/obligations.ml +++ b/toplevel/obligations.ml @@ -502,14 +502,15 @@ let subst_body expand prg = let declare_definition prg = let body, typ = subst_body true prg in let ce = - { const_entry_body = body; + { const_entry_body = Future.from_val (body,Declareops.no_seff); const_entry_secctx = None; const_entry_type = Some typ; const_entry_opaque = false; const_entry_inline_code = false} in progmap_remove prg; - !declare_definition_ref prg.prg_name prg.prg_kind ce prg.prg_implicits + !declare_definition_ref prg.prg_name + prg.prg_kind ce prg.prg_implicits (fun local gr -> prg.prg_hook local gr; gr) open Pp @@ -550,27 +551,35 @@ let declare_mutual_definition l = let fixkind = Option.get first.prg_fixkind in let arrrec, recvec = Array.of_list fixtypes, Array.of_list fixdefs in let fixdecls = (Array.of_list (List.map (fun x -> Name x.prg_name) l), arrrec, recvec) in - let (local,kind) = first.prg_kind in let fixnames = first.prg_deps in - let kind = if fixkind != IsCoFixpoint then Fixpoint else CoFixpoint in + let kind = + fst first.prg_kind, + if fixkind != IsCoFixpoint then Fixpoint else CoFixpoint in let indexes, fixdecls = match fixkind with | IsFixpoint wfl -> let possible_indexes = - List.map3 compute_possible_guardness_evidences wfl fixdefs fixtypes in - let indexes = Pretyping.search_guard Loc.ghost (Global.env ()) possible_indexes fixdecls in - Some indexes, List.map_i (fun i _ -> mkFix ((indexes,i),fixdecls)) 0 l + List.map3 compute_possible_guardness_evidences + wfl fixdefs fixtypes in + let indexes = + Pretyping.search_guard Loc.ghost (Global.env()) + possible_indexes fixdecls in + Some indexes, + List.map_i (fun i _ -> + mkFix ((indexes,i),fixdecls),Declareops.no_seff) 0 l | IsCoFixpoint -> - None, List.map_i (fun i _ -> mkCoFix (i,fixdecls)) 0 l + None, + List.map_i (fun i _ -> + mkCoFix (i,fixdecls),Declareops.no_seff) 0 l in (* Declare the recursive definitions *) - let kns = List.map4 (!declare_fix_ref (local, kind)) fixnames fixdecls fixtypes fiximps in + let kns = List.map4 (!declare_fix_ref kind) fixnames fixdecls fixtypes fiximps in (* Declare notations *) List.iter Metasyntax.add_notation_interpretation first.prg_notations; Declare.recursive_message (fixkind != IsCoFixpoint) indexes fixnames; let gr = List.hd kns in let kn = match gr with ConstRef kn -> kn | _ -> assert false in - first.prg_hook local gr; + first.prg_hook (fst first.prg_kind) gr; List.iter progmap_remove l; kn let shrink_body c = @@ -594,7 +603,7 @@ let declare_obligation prg obl body = if get_shrink_obligations () then shrink_body body else [], body, [||] in let ce = - { const_entry_body = body; + { const_entry_body = Future.from_val(body,Declareops.no_seff); const_entry_secctx = None; const_entry_type = if ctx = [] then Some ty else None; const_entry_opaque = opaque; @@ -610,7 +619,7 @@ let declare_obligation prg obl body = definition_message obl.obl_name; { obl with obl_body = Some (it_mkLambda_or_LetIn (mkApp (mkConst constant, args)) ctx) } -let init_prog_info n b t deps fixkind notations obls impls kind reduce hook = +let init_prog_info n b t deps fixkind notations obls impls k reduce hook = let obls', b = match b with | None -> @@ -632,7 +641,8 @@ let init_prog_info n b t deps fixkind notations obls impls kind reduce hook = { prg_name = n ; prg_body = b; prg_type = reduce t; prg_obligations = (obls', Array.length obls'); prg_deps = deps; prg_fixkind = fixkind ; prg_notations = notations ; - prg_implicits = impls; prg_kind = kind; prg_reduce = reduce; prg_hook = hook; } + prg_implicits = impls; prg_kind = k; prg_reduce = reduce; + prg_hook = hook; } let get_prog name = let prg_infos = !from_prg in @@ -645,7 +655,11 @@ let get_prog name = match n with 0 -> raise (NoObligations None) | 1 -> map_first prg_infos - | _ -> error "More than one program with unsolved obligations") + | _ -> + error ("More than one program with unsolved obligations: "^ + String.concat ", " + (List.map string_of_id + (ProgMap.fold (fun k _ s -> k::s) prg_infos [])))) let get_any_prog () = let prg_infos = !from_prg in @@ -737,19 +751,14 @@ let rec string_of_list sep f = function (* Solve an obligation using tactics, return the corresponding proof term *) let solve_by_tac evi t = let id = Id.of_string "H" in - try - Pfedit.start_proof id goal_kind evi.evar_hyps evi.evar_concl - (fun _ _ -> ()); - Pfedit.by (tclCOMPLETE t); - let _,(const,_,_,_) = Pfedit.cook_proof ignore in - Pfedit.delete_current_proof (); - Inductiveops.control_only_guard (Global.env ()) - const.Entries.const_entry_body; - const.Entries.const_entry_body - with reraise -> - let reraise = Errors.push reraise in - Pfedit.delete_current_proof(); - raise reraise + let entry = Pfedit.build_constant_by_tactic + id ~goal_kind evi.evar_hyps evi.evar_concl (tclCOMPLETE t) in + let env = Global.env () in + let entry = Term_typing.handle_side_effects env entry in + let body, eff = Future.force entry.Entries.const_entry_body in + assert(eff = Declareops.no_seff); + Inductiveops.control_only_guard (Global.env ()) body; + body let rec solve_obligation prg num tac = let user_num = succ num in diff --git a/toplevel/obligations.mli b/toplevel/obligations.mli index bc092a1ce..ddc99a96c 100644 --- a/toplevel/obligations.mli +++ b/toplevel/obligations.mli @@ -22,7 +22,7 @@ open Tacexpr (** Forward declaration. *) val declare_fix_ref : (definition_kind -> Id.t -> - constr -> types -> Impargs.manual_implicits -> global_reference) ref + Entries.proof_output -> types -> Impargs.manual_implicits -> global_reference) ref val declare_definition_ref : (Id.t -> definition_kind -> diff --git a/toplevel/record.ml b/toplevel/record.ml index ff918293b..476da3d0e 100644 --- a/toplevel/record.ml +++ b/toplevel/record.ml @@ -189,7 +189,8 @@ let declare_projections indsp ?(kind=StructureComponent) ?name coers fieldimpls let kn = try let cie = { - const_entry_body = proj; + const_entry_body = + Future.from_val (proj,Declareops.no_seff); const_entry_secctx = None; const_entry_type = Some projtyp; const_entry_opaque = false; @@ -293,7 +294,8 @@ let declare_class finite def infer id idbuild paramimpls params arity fieldimpls let class_body = it_mkLambda_or_LetIn field params in let class_type = Option.map (fun ar -> it_mkProd_or_LetIn ar params) arity in let class_entry = - { const_entry_body = class_body; + { const_entry_body = + Future.from_val (class_body,Declareops.no_seff); const_entry_secctx = None; const_entry_type = class_type; const_entry_opaque = false; @@ -306,7 +308,8 @@ let declare_class finite def infer id idbuild paramimpls params arity fieldimpls let proj_type = it_mkProd_or_LetIn (mkProd(Name (snd id), inst_type, lift 1 field)) params in let proj_body = it_mkLambda_or_LetIn (mkLambda (Name (snd id), inst_type, mkRel 1)) params in let proj_entry = - { const_entry_body = proj_body; + { const_entry_body = + Future.from_val (proj_body,Declareops.no_seff); const_entry_secctx = None; const_entry_type = Some proj_type; const_entry_opaque = false; diff --git a/toplevel/stm.ml b/toplevel/stm.ml new file mode 100644 index 000000000..004ac9428 --- /dev/null +++ b/toplevel/stm.ml @@ -0,0 +1,952 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* true | _ -> false in + if internal_command expr then () + else begin + Pp.set_id_for_feedback (Interface.State id); + try Vernacentries.interp ~verbosely ?proof (loc, expr) + with e -> raise (Errors.push (Cerrors.process_vernac_interp_error e)) + end + +type ast = bool * Loc.t * vernac_expr +let pr_ast (_, _, v) = pr_vernac v + +module Vcs_ = Vcs.Make(StateidOrderedType) + +type branch_type = [ `Master | `Proof of string * int ] +type cmd_t = ast +type fork_t = ast * Vcs_.branch_name * Names.Id.t list +type qed_t = + ast * vernac_qed_type * (Vcs_.branch_name * branch_type Vcs_.branch_info) +type seff_t = ast option +type alias_t = state_id +type transaction = + | Cmd of cmd_t + | Fork of fork_t + | Qed of qed_t + | Sideff of seff_t + | Alias of alias_t + | Noop +type step = + [ `Cmd of cmd_t + | `Fork of fork_t + | `Qed of qed_t * state_id + | `Sideff of [ `Ast of ast * state_id | `Id of state_id ] + | `Alias of alias_t ] +type visit = { step : step; next : state_id } + +type state = States.state * Proof_global.state +type state_info = { (* Make private *) + mutable n_reached : int; + mutable n_goals : int; + mutable state : state option; +} +let default_info () = { n_reached = 0; n_goals = 0; state = None } + +(* Functions that work on a Vcs with a specific branch type *) +module Vcs_aux : sig + + val proof_nesting : (branch_type, 't,'i) Vcs_.t -> int + val find_proof_at_depth : + (branch_type, 't, 'i) Vcs_.t -> int -> + Vcs_.branch_name * branch_type Vcs_.branch_info + +end = struct (* {{{ *) + + let proof_nesting vcs = + List.fold_left max 0 + (List.map_filter + (function { Vcs_.kind = `Proof (_,n) } -> Some n | _ -> None) + (List.map (Vcs_.get_branch vcs) (Vcs_.branches vcs))) + + let find_proof_at_depth vcs pl = + try List.find (function + | _, { Vcs_.kind = `Proof(m, n) } -> n = pl + | _ -> false) + (List.map (fun h -> h, Vcs_.get_branch vcs h) (Vcs_.branches vcs)) + with Not_found -> failwith "find_proof_at_depth" + +end (* }}} *) + +(* Imperative wrap around VCS to obtain _the_ VCS *) +module VCS : sig + + type id = state_id + type branch_name = Vcs_.branch_name + type 'branch_type branch_info = 'branch_type Vcs_.branch_info = { + kind : [> `Master] as 'branch_type; + root : id; + pos : id; + } + + type vcs = (branch_type, transaction, state_info) Vcs_.t + + val init : id -> unit + + val string_of_branch_name : branch_name -> string + + val current_branch : unit -> branch_name + val checkout : branch_name -> unit + val master : branch_name + val branches : unit -> branch_name list + val get_branch : branch_name -> branch_type branch_info + val get_branch_pos : branch_name -> id + val new_node : unit -> id + val merge : id -> ours:transaction -> ?into:branch_name -> branch_name -> unit + val delete_branch : branch_name -> unit + val commit : id -> transaction -> unit + val mk_branch_name : ast -> branch_name + val branch : branch_name -> branch_type -> unit + + val get_info : id -> state_info + val reached : id -> bool -> unit + val goals : id -> int -> unit + val set_state : id -> state -> unit + val forget_state : id -> unit + + val create_cluster : id list -> unit + + val proof_nesting : unit -> int + val checkout_shallowest_proof_branch : unit -> unit + val propagate_sideff : ast option -> unit + + val visit : id -> visit + + val print : unit -> unit + + val backup : unit -> vcs + val restore : vcs -> unit + +end = struct (* {{{ *) + + include Vcs_ + + module StateidSet = Set.Make(StateidOrderedType) + open Printf + + let print_dag vcs () = (* {{{ *) + let fname = "stm" in + let string_of_transaction = function + | Cmd t | Fork (t, _, _) -> + (try string_of_ppcmds (pr_ast t) with _ -> "ERR") + | Sideff (Some t) -> + sprintf "Sideff(%s)" + (try string_of_ppcmds (pr_ast t) with _ -> "ERR") + | Sideff None -> "EnvChange" + | Noop -> " " + | Alias id -> sprintf "Alias(%s)" (string_of_state_id id) + | Qed (a,_,_) -> string_of_ppcmds (pr_ast a) in + let is_green id = + match get_info vcs id with + | Some { state = Some _ } -> true + | _ -> false in + let is_red id = + match get_info vcs id with + | Some s -> s.n_reached = ~-1 + | _ -> false in + let head = current_branch vcs in + let heads = + List.map (fun x -> x, (get_branch vcs x).pos) (branches vcs) in + let graph = dag vcs in + let quote s = + Str.global_replace (Str.regexp "\n") "
" + (Str.global_replace (Str.regexp "<") "<" + (Str.global_replace (Str.regexp ">") ">" + (Str.global_replace (Str.regexp "\"") """ + (Str.global_replace (Str.regexp "&") "&" + (String.sub s 0 (min (String.length s) 20)))))) in + let fname_dot, fname_ps = + let f = "/tmp/" ^ Filename.basename fname in + f ^ ".dot", f ^ ".pdf" in + let node id = "s" ^ string_of_state_id id in + let edge tr = + sprintf "<%s>" + (quote (string_of_transaction tr)) in + let ids = ref StateidSet.empty in + let clus = Hashtbl.create 13 in + let node_info id = + match get_info vcs id with + | None -> "" + | Some info -> + sprintf "<%s" (string_of_state_id id) ^ + sprintf " r:%d g:%d>" + info.n_reached info.n_goals in + let color id = + if is_red id then "red" else if is_green id then "green" else "white" in + let nodefmt oc id = + fprintf oc "%s [label=%s,style=filled,fillcolor=%s];\n" + (node id) (node_info id) (color id) in + let add_to_clus_or_ids from cf = + match cf with + | None -> ids := StateidSet.add from !ids; false + | Some c -> Hashtbl.replace clus c + (try let n = Hashtbl.find clus c in from::n + with Not_found -> [from]); true in + let oc = open_out fname_dot in + output_string oc "digraph states {\nsplines=ortho\n"; + Dag.iter graph (fun from cf _ l -> + let c1 = add_to_clus_or_ids from cf in + List.iter (fun (dest, trans) -> + let c2 = add_to_clus_or_ids dest (Dag.cluster_of graph dest) in + fprintf oc "%s -> %s [label=%s,labelfloat=%b];\n" + (node from) (node dest) (edge trans) (c1 && c2)) l + ); + StateidSet.iter (nodefmt oc) !ids; + Hashtbl.iter (fun c nodes -> + fprintf oc "subgraph cluster_%s {\n" (Dag.string_of_cluster_id c); + List.iter (nodefmt oc) nodes; + fprintf oc "color=blue; }\n" + ) clus; + List.iteri (fun i (b,id) -> + let shape = if head = b then "box3d" else "box" in + fprintf oc "b%d -> %s;\n" i (node id); + fprintf oc "b%d [shape=%s,label=\"%s\"];\n" i shape + (string_of_branch_name b); + ) heads; + output_string oc "}\n"; + close_out oc; + ignore(Sys.command + ("dot -Tpdf -Gcharset=latin1 " ^ fname_dot ^ " -o" ^ fname_ps)) + (* }}} *) + + type vcs = (branch_type, transaction, state_info) t + let vcs : vcs ref = ref (empty dummy_state_id) + + let init id = + vcs := empty id; + vcs := set_info !vcs id (default_info ()) + + let current_branch () = current_branch !vcs + + let checkout head = vcs := checkout !vcs head + let master = master + let branches () = branches !vcs + let get_branch head = get_branch !vcs head + let get_branch_pos head = (get_branch head).pos + let new_node () = + let id = Stateid.fresh_state_id () in + vcs := set_info !vcs id (default_info ()); + id + let merge id ~ours ?into branch = + vcs := merge !vcs id ~ours ~theirs:Noop ?into branch + let delete_branch branch = vcs := delete_branch !vcs branch + let commit id t = vcs := commit !vcs id t + let mk_branch_name (_, _, x) = mk_branch_name + (match x with + | VernacDefinition (_,(_,i),_) -> string_of_id i + | VernacStartTheoremProof (_,[Some (_,i),_],_) -> string_of_id i + | _ -> "branch") + let branch name kind = vcs := branch !vcs name kind + let get_info id = + match get_info !vcs id with + | Some x -> x + | None -> assert false + let set_state id s = (get_info id).state <- Some s + let forget_state id = (get_info id).state <- None + let reached id b = + let info = get_info id in + if b then info.n_reached <- info.n_reached + 1 + else info.n_reached <- -1 + let goals id n = (get_info id).n_goals <- n + let create_cluster l = vcs := create_cluster !vcs l + + let proof_nesting () = Vcs_aux.proof_nesting !vcs + + let checkout_shallowest_proof_branch () = + let pl = proof_nesting () in + try + let branch, mode = match Vcs_aux.find_proof_at_depth !vcs pl with + | h, { Vcs_.kind = `Proof (m, _) } -> h, m | _ -> assert false in + checkout branch; + Proof_global.activate_proof_mode mode + with Failure _ -> + checkout master; + Proof_global.disactivate_proof_mode "Classic" (* XXX *) + + (* copies the transaction on every open branch *) + let propagate_sideff t = + List.iter (fun b -> + checkout b; + let id = new_node () in + merge id ~ours:(Sideff t) ~into:b master) + (List.filter ((<>) master) (branches ())) + + let visit id = + match Dag.from_node (dag !vcs) id with + | [n, Cmd x] -> { step = `Cmd x; next = n } + | [n, Alias x] -> { step = `Alias x; next = n } + | [n, Fork x] -> { step = `Fork x; next = n } + | [n, Qed x; p, Noop] + | [p, Noop; n, Qed x] -> { step = `Qed (x,p); next = n } + | [n, Sideff None; p, Noop] + | [p, Noop; n, Sideff None]-> { step = `Sideff (`Id p); next = n } + | [n, Sideff (Some x); p, Noop] + | [p, Noop; n, Sideff (Some x)]-> { step = `Sideff (`Ast (x,p)); next = n } + | _ -> anomaly (str "Malformed VCS, or visiting the root") + + module NB : sig + + val command : (unit -> unit) -> unit + + end = struct + + let m = Mutex.create () + let c = Condition.create () + let job = ref None + let worker = ref None + + let set_last_job j = + Mutex.lock m; + job := Some j; + Condition.signal c; + Mutex.unlock m + + let get_last_job () = + Mutex.lock m; + while !job = None do Condition.wait c m; done; + match !job with + | None -> assert false + | Some x -> job := None; Mutex.unlock m; x + + let run_command () = + while true do get_last_job () () done + + let command job = + set_last_job job; + if !worker = None then worker := Some (Thread.create run_command ()) + + end + + let print () = + if not !Flags.debug then () else NB.command (print_dag !vcs) + + let backup () = !vcs + let restore v = vcs := v + +end (* }}} *) + +(* Fills in the nodes of the VCS *) +module State : sig + + (** The function is from unit, so it uses the current state to define + a new one. I.e. one may been to install the right state before + defining a new one. + + Warning: an optimization requires that state modifying functions + are always executed using this wrapper. *) + val define : ?cache:bool -> (unit -> unit) -> state_id -> unit + + val install_cached : state_id -> unit + val is_cached : state_id -> bool + + val exn_on : state_id -> ?valid:state_id -> exn -> exn + +end = struct (* {{{ *) + + (* cur_id holds dummy_state_id in case the last attempt to define a state + * failed, so the global state may contain garbage *) + let cur_id = ref dummy_state_id + + (* helpers *) + let freeze_global_state () = + States.freeze ~marshallable:false, Proof_global.freeze () + let unfreeze_global_state (s,p) = + States.unfreeze s; Proof_global.unfreeze p + + (* hack to make futures functional *) + let in_t, out_t = Dyn.create "state4future" + let () = Future.set_freeze + (fun () -> in_t (freeze_global_state (), !cur_id)) + (fun t -> let s,i = out_t t in unfreeze_global_state s; cur_id := i) + + let is_cached id = + id = !cur_id || + match VCS.get_info id with + | { state = Some _ } -> true + | _ -> false + + let install_cached id = + if id = !cur_id then () else (* optimization *) + let s = + match VCS.get_info id with + | { state = Some s } -> s + | _ -> anomaly (str "unfreezing a non existing state") in + unfreeze_global_state s; cur_id := id + + let freeze id = VCS.set_state id (freeze_global_state ()) + + let exn_on id ?valid e = + Stateid.add_state_id e ?valid id + + let define ?(cache=false) f id = + if is_cached id then + anomaly (str"defining state "++str(string_of_state_id id)++str" twice"); + try + prerr_endline ("defining " ^ + string_of_state_id id ^ " (cache=" ^ string_of_bool cache ^ ")"); + f (); + if cache then freeze id; + cur_id := id; + feedback ~state_id:id Interface.Processed; + VCS.reached id true; + if Proof_global.there_are_pending_proofs () then + VCS.goals id (Proof_global.get_open_goals ()); + with e -> + let e = Errors.push e in + let good_id = !cur_id in + cur_id := dummy_state_id; + VCS.reached id false; + match Stateid.get_state_id e with + | Some _ -> raise e + | None -> raise (exn_on id ~valid:good_id e) + +end +(* }}} *) + + +(* Runs all transactions needed to reach a state *) +module Reach : sig + + val known_state : cache:bool -> state_id -> unit + +end = struct (* {{{ *) + +let pstate = ["meta counter"; "evar counter"; "program-tcc-table"] + +let collect_proof cur hd id = + let rec collect last accn id = + let view = VCS.visit id in + match last, view.step with + | _, `Cmd x -> collect (Some (id,x)) (id::accn) view.next + | _, `Alias _ -> collect None (id::accn) view.next + | Some (parent, (_,_,VernacExactProof _)), `Fork _ -> + `NotOptimizable `Immediate + | Some (parent, (_,_,VernacProof(_,Some _) as v)), `Fork (_, hd', _) -> + assert( hd = hd' ); + `Optimizable (parent, Some v, accn) + | Some (parent, _), `Fork (_, hd', _) -> + assert( hd = hd' ); + `MaybeOptimizable (parent, accn) + | _, `Sideff se -> collect None (id::accn) view.next + | _ -> `NotOptimizable `Unknown in + if State.is_cached id then `NotOptimizable `Unknown + else collect (Some cur) [] id + +let known_state ~cache id = + + (* ugly functions to process nested lemmas, i.e. hard to reproduce + * side effects *) + let cherry_pick_non_pstate () = + Summary.freeze_summary ~marshallable:false ~complement:true pstate, + Lib.freeze ~marshallable:false in + let inject_non_pstate (s,l) = Summary.unfreeze_summary s; Lib.unfreeze l in + + let rec pure_cherry_pick_non_pstate id = Future.purify (fun id -> + prerr_endline ("cherry-pick non pstate " ^ string_of_state_id id); + reach id; + cherry_pick_non_pstate ()) id + + (* traverses the dag backward from nodes being already calculated *) + and reach ?(cache=cache) id = + prerr_endline ("reaching: " ^ string_of_state_id id); + if State.is_cached id then begin + State.install_cached id; + feedback ~state_id:id Interface.Processed; + prerr_endline ("reached (cache)") + end else + let step, cache_step = + let view = VCS.visit id in + match view.step with + | `Alias id -> + (fun () -> + reach view.next; reach id; Vernacentries.try_print_subgoals ()), + cache + | `Cmd (x,_) -> (fun () -> reach view.next; interp id x), cache + | `Fork (x,_,_) -> (fun () -> reach view.next; interp id x), true + | `Qed ((x,keep,(hd,_)), eop) -> + let rec aux = function + | `Optimizable (start, proof_using_ast, nodes) -> + (fun () -> + prerr_endline ("Optimizable " ^ string_of_state_id id); + VCS.create_cluster nodes; + begin match keep with + | KeepProof -> + let f = Future.create (fun () -> reach eop) in + reach start; + let proof = + Proof_global.close_future_proof + ~fix_exn:(State.exn_on id ~valid:eop) f in + reach view.next; + interp id ~proof x; + | DropProof -> + reach view.next; + Option.iter (interp start) proof_using_ast; + interp id x + end; + Proof_global.discard_all ()) + | `NotOptimizable `Immediate -> assert (view.next = eop); + (fun () -> reach eop; interp id x; Proof_global.discard_all ()) + | `NotOptimizable `Unknown -> + (fun () -> + prerr_endline ("NotOptimizable " ^ string_of_state_id id); + reach eop; + begin match keep with + | KeepProof -> + let proof = Proof_global.close_proof () in + reach view.next; + interp id ~proof x + | DropProof -> + reach view.next; + interp id x + end; + Proof_global.discard_all ()) + | `MaybeOptimizable (start, nodes) -> + (fun () -> + prerr_endline ("MaybeOptimizable " ^ string_of_state_id id); + reach ~cache:true start; + (* no sections and no modules *) + if Environ.named_context (Global.env ()) = [] && + Safe_typing.is_curmod_library (Global.safe_env ()) then + aux (`Optimizable (start, None, nodes)) () + else + aux (`NotOptimizable `Unknown) ()) + in + aux (collect_proof (view.next, x) hd eop), true + | `Sideff (`Ast (x,_)) -> + (fun () -> reach view.next; interp id x), cache + | `Sideff (`Id origin) -> + (fun () -> + let s = pure_cherry_pick_non_pstate origin in + reach view.next; + inject_non_pstate s), + cache + in + State.define ~cache:cache_step step id; + prerr_endline ("reached: "^ string_of_state_id id) in + reach id + +end (* }}} *) + +(* The backtrack module simulates the classic behavior of a linear document *) +module Backtrack : sig + + val record : unit -> unit + val backto : state_id -> unit + val cur : unit -> state_id + + (* we could navigate the dag, but this ways easy *) + val branches_of : state_id -> VCS.branch_name list + + (* To be installed during initialization *) + val undo_vernac_classifier : vernac_expr -> vernac_classification + +end = struct (* {{{ *) + + module S = Searchstack + + type hystory_elt = { + id : state_id ; + vcs : VCS.vcs; + label : Names.Id.t list; (* To implement a limited form of reset *) + } + + let history : hystory_elt S.t = S.create () + + let cur () = + if S.is_empty history then anomaly (str "Empty history"); + (S.top history).id + + let record () = + let id = VCS.get_branch_pos (VCS.current_branch ()) in + S.push { + id = id; + vcs = VCS.backup (); + label = + if id = initial_state_id || id = dummy_state_id then [] else + match VCS.visit id with + | { step = `Fork (_,_,l) } -> l + | { step = `Cmd (_,_, VernacFixpoint (_,l)) } -> + List.map (fun (((_,id),_,_,_,_),_) -> id) l + | { step = `Cmd (_,_, VernacCoFixpoint (_,l)) } -> + List.map (fun (((_,id),_,_,_),_) -> id) l + | { step = `Cmd (_,_, VernacAssumption (_,_,l)) } -> + List.flatten (List.map (fun (_,(lid,_)) -> List.map snd lid) l) + | { step = `Cmd (_,_, VernacInductive (_,_,l)) } -> + List.map (fun (((_,(_,id)),_,_,_,_),_) -> id) l + | { step = `Cmd (_,_, (VernacDefinition (_,(_,id),DefineBody _) | + VernacDeclareModuleType ((_,id),_,_,_) | + VernacDeclareModule (_,(_,id),_,_) | + VernacDefineModule (_,(_,id),_,_,_))) } -> [id] + | _ -> [] } + history + + let backto oid = + assert(not (S.is_empty history)); + let rec pop_until_oid () = + let id = (S.top history).id in + if id = initial_state_id || id = oid then () + else begin ignore (S.pop history); pop_until_oid (); end in + pop_until_oid (); + let backup = S.top history in + VCS.restore backup.vcs; + if backup.id <> oid then anomaly (str "Backto an unknown state") + + let branches_of id = + try + let s = S.find (fun n s -> + if s.id = id then `Stop s else `Cont ()) () history in + Vcs_.branches s.vcs + with Not_found -> assert false + + let undo_vernac_classifier = function + | VernacResetInitial -> + VtStm (VtBack initial_state_id, true), VtNow + | VernacResetName (_,name) -> + msg_warning + (str"Reset not implemented for automatically generated constants"); + (try + let s = + S.find (fun b s -> + if b then `Stop s else `Cont (List.mem name s.label)) + false history in + VtStm (VtBack s.id, true), VtNow + with Not_found -> + VtStm (VtBack (S.top history).id, true), VtNow) + | VernacBack n -> + let s = S.find (fun n s -> + if n = 0 then `Stop s else `Cont (n-1)) n history in + VtStm (VtBack s.id, true), VtNow + | VernacUndo n -> + let s = S.find (fun n s -> + if n = 0 then `Stop s else `Cont (n-1)) n history in + VtStm (VtBack s.id, true), VtLater + | VernacUndoTo _ + | VernacRestart as e -> + let m = match e with VernacUndoTo m -> m | _ -> 0 in + let vcs = (S.top history).vcs in + let cb, _ = + Vcs_aux.find_proof_at_depth vcs (Vcs_aux.proof_nesting vcs) in + let n = S.find (fun n { vcs } -> + if List.mem cb (Vcs_.branches vcs) then `Cont (n+1) else `Stop n) + 0 history in + let s = S.find (fun n s -> + if n = 0 then `Stop s else `Cont (n-1)) (n-m-1) history in + VtStm (VtBack s.id, true), VtLater + | VernacAbortAll -> + let s = S.find (fun () s -> + if List.length (Vcs_.branches s.vcs) = 1 then `Stop s else `Cont ()) + () history in + VtStm (VtBack s.id, true), VtLater + | VernacBacktrack (id,_,_) + | VernacBackTo id -> + VtStm (VtBack (Stateid.state_id_of_int id), true), VtNow + | _ -> VtUnknown, VtNow + +end (* }}} *) + +let init () = + VCS.init initial_state_id; + declare_vernac_classifier "Stm" Backtrack.undo_vernac_classifier; + State.define ~cache:true (fun () -> ()) initial_state_id; + Backtrack.record () + +let observe id = + let vcs = VCS.backup () in + try + Reach.known_state ~cache:(interactive ()) id; + VCS.print () + with e -> + let e = Errors.push e in + VCS.print (); + VCS.restore vcs; + raise e + +let finish () = + observe (VCS.get_branch_pos (VCS.current_branch ())); + VCS.print () + +let join_aborted_proofs () = + let rec aux id = + if id = initial_state_id then () else + let view = VCS.visit id in + match view.step with + | `Qed ((_,DropProof,_),eop) -> observe eop; aux view.next + | `Sideff _ | `Alias _ | `Cmd _ | `Fork _ | `Qed _ -> aux view.next + in + aux (VCS.get_branch_pos VCS.master) + +let join () = + finish (); + VCS.print (); + prerr_endline "Joining the environment"; + Global.join_safe_environment (); + VCS.print (); + prerr_endline "Joining the aborted proofs"; + join_aborted_proofs (); + VCS.print () + +let merge_proof_branch x keep branch = + if branch = VCS.master then + raise(State.exn_on dummy_state_id Proof_global.NoCurrentProof); + let info = VCS.get_branch branch in + VCS.checkout VCS.master; + let id = VCS.new_node () in + VCS.merge id ~ours:(Qed (x,keep,(branch, info))) branch; + VCS.delete_branch branch; + if keep = KeepProof then VCS.propagate_sideff None + +let process_transaction verbosely (loc, expr) = + let warn_if_pos a b = + if b then msg_warning(pr_ast a ++ str" should not be part of a script") in + let v, x = expr, (verbosely && Flags.is_verbose(), loc, expr) in + prerr_endline ("{{{ execute: "^ string_of_ppcmds (pr_ast x)); + let vcs = VCS.backup () in + try + let head = VCS.current_branch () in + VCS.checkout head; + begin + let c = classify_vernac v in + prerr_endline (" classified as: " ^ string_of_vernac_classification c); + match c with + (* Joining various parts of the document *) + | VtStm (VtJoinDocument, b), VtNow -> warn_if_pos x b; join () + | VtStm (VtFinish, b), VtNow -> warn_if_pos x b; finish () + | VtStm (VtObserve id, b), VtNow -> warn_if_pos x b; observe id + | VtStm ((VtObserve _ | VtFinish | VtJoinDocument), _), VtLater -> + anomaly(str"classifier: join actions cannot be classified as VtLater") + + (* Back *) + | VtStm (VtBack oid, true), w -> + let id = VCS.new_node () in + let bl = Backtrack.branches_of oid in + List.iter (fun branch -> + if not (List.mem branch bl) then + merge_proof_branch + (false,Loc.ghost,VernacAbortAll) DropProof branch) + (VCS.branches ()); + VCS.checkout_shallowest_proof_branch (); + VCS.commit id (Alias oid); + Backtrack.record (); if w = VtNow then finish () + | VtStm (VtBack id, false), VtNow -> + prerr_endline ("undo to state " ^ string_of_state_id id); + Backtrack.backto id; + VCS.checkout_shallowest_proof_branch (); + Reach.known_state ~cache:(interactive ()) id; + | VtStm (VtBack id, false), VtLater -> + anomaly(str"classifier: VtBack + VtLater must imply part_of_script") + + (* Query *) + | VtQuery false, VtNow -> + finish (); + (try Future.purify (interp dummy_state_id) (true,loc,expr) + with e when Errors.noncritical e -> + let e = Errors.push e in + raise(State.exn_on dummy_state_id e)) + | VtQuery true, w -> + let id = VCS.new_node () in + VCS.commit id (Cmd x); + Backtrack.record (); if w = VtNow then finish () + | VtQuery false, VtLater -> + anomaly(str"classifier: VtQuery + VtLater must imply part_of_script") + + (* Proof *) + | VtStartProof (mode, names), w -> + let id = VCS.new_node () in + let bname = VCS.mk_branch_name x in + VCS.checkout VCS.master; + VCS.commit id (Fork (x, bname, names)); + VCS.branch bname (`Proof (mode, VCS.proof_nesting () + 1)); + Proof_global.activate_proof_mode mode; + Backtrack.record (); if w = VtNow then finish () + | VtProofStep, w -> + let id = VCS.new_node () in + VCS.commit id (Cmd x); + Backtrack.record (); if w = VtNow then finish () + | VtQed keep, w -> + merge_proof_branch x keep head; + VCS.checkout_shallowest_proof_branch (); + Backtrack.record (); if w = VtNow then finish () + + (* Side effect on all branches *) + | VtSideff, w -> + let id = VCS.new_node () in + VCS.checkout VCS.master; + VCS.commit id (Cmd x); + VCS.propagate_sideff (Some x); + VCS.checkout_shallowest_proof_branch (); + Backtrack.record (); if w = VtNow then finish () + + (* Unknown: we execute it, check for open goals and propagate sideeff *) + | VtUnknown, VtNow -> + let id = VCS.new_node () in + let step () = + VCS.checkout VCS.master; + let mid = VCS.get_branch_pos VCS.master in + Reach.known_state ~cache:(interactive ()) mid; + interp id x; + (* Vernac x may or may not start a proof *) + if head = VCS.master && + Proof_global.there_are_pending_proofs () + then begin + let bname = VCS.mk_branch_name x in + VCS.commit id (Fork (x,bname,[])); + VCS.branch bname (`Proof ("Classic", VCS.proof_nesting () + 1)) + end else begin + VCS.commit id (Cmd x); + VCS.propagate_sideff (Some x); + VCS.checkout_shallowest_proof_branch (); + end in + State.define ~cache:true step id; + Backtrack.record () + + | VtUnknown, VtLater -> + anomaly(str"classifier: VtUnknown must imply VtNow") + end; + prerr_endline "executed }}}"; + VCS.print () + with e -> + match Stateid.get_state_id e with + | None -> + VCS.restore vcs; + VCS.print (); + anomaly (str ("execute: "^ + string_of_ppcmds (Errors.print_no_report e) ^ "}}}")) + | Some (_, id) -> + let e = Errors.push e in + prerr_endline ("Failed at state " ^ Stateid.string_of_state_id id); + VCS.restore vcs; + VCS.print (); + raise e + +(* Query API *) + +let get_current_state () = Backtrack.cur () + +let current_proof_depth () = + let head = VCS.current_branch () in + match VCS.get_branch head with + | { VCS.kind = `Master } -> 0 + | { VCS.pos = cur; VCS.kind = `Proof (_,n); VCS.root = root } -> + let rec distance cur = + if cur = root then 0 + else 1 + distance (VCS.visit cur).next in + distance cur + +let unmangle n = + let n = VCS.string_of_branch_name n in + let idx = String.index n '_' + 1 in + Names.id_of_string (String.sub n idx (String.length n - idx)) + +let proofname b = match VCS.get_branch b with + | { VCS.kind = `Proof _ } -> Some b + | _ -> None + +let get_all_proof_names () = + List.map unmangle (List.map_filter proofname (VCS.branches ())) + +let get_current_proof_name () = + Option.map unmangle (proofname (VCS.current_branch ())) + +let get_script prf = + let rec find acc id = + if id = initial_state_id || id = dummy_state_id then acc else + let view = VCS.visit id in + match view.step with + | `Fork (_,_,ns) when List.mem prf ns -> acc + | `Qed ((x,_,_), proof) -> find [pi3 x, (VCS.get_info id).n_goals] proof + | `Sideff (`Ast (x,id)) -> find ((pi3 x, (VCS.get_info id).n_goals)::acc) id + | `Sideff (`Id id) -> find acc id + | `Cmd x -> find ((pi3 x, (VCS.get_info id).n_goals)::acc) view.next + | `Alias id -> find acc id + | `Fork _ -> find acc view.next + in + find [] (VCS.get_branch_pos VCS.master) + +(* indentation code for Show Script, initially contributed + by D. de Rauglaudre *) + +let indent_script_item ((ng1,ngl1),nl,beginend,ppl) (cmd,ng) = + (* ng1 : number of goals remaining at the current level (before cmd) + ngl1 : stack of previous levels with their remaining goals + ng : number of goals after the execution of cmd + beginend : special indentation stack for { } *) + let ngprev = List.fold_left (+) ng1 ngl1 in + let new_ngl = + if ng > ngprev then + (* We've branched *) + (ng - ngprev + 1, ng1 - 1 :: ngl1) + else if ng < ngprev then + (* A subgoal have been solved. Let's compute the new current level + by discarding all levels with 0 remaining goals. *) + let _ = assert (Int.equal ng (ngprev - 1)) in + let rec loop = function + | (0, ng2::ngl2) -> loop (ng2,ngl2) + | p -> p + in loop (ng1-1, ngl1) + else + (* Standard case, same goal number as before *) + (ng1, ngl1) + in + (* When a subgoal have been solved, separate this block by an empty line *) + let new_nl = (ng < ngprev) + in + (* Indentation depth *) + let ind = List.length ngl1 + in + (* Some special handling of bullets and { }, to get a nicer display *) + let pred n = max 0 (n-1) in + let ind, nl, new_beginend = match cmd with + | VernacSubproof _ -> pred ind, nl, (pred ind)::beginend + | VernacEndSubproof -> List.hd beginend, false, List.tl beginend + | VernacBullet _ -> pred ind, nl, beginend + | _ -> ind, nl, beginend + in + let pp = + (if nl then fnl () else mt ()) ++ + (hov (ind+1) (str (String.make ind ' ') ++ Ppvernac.pr_vernac cmd)) + in + (new_ngl, new_nl, new_beginend, pp :: ppl) + +let show_script ?proof () = + try + let prf = + match proof with + | None -> Pfedit.get_current_proof_name () + | Some (id,_) -> id in + let cmds = get_script prf in + let _,_,_,indented_cmds = + List.fold_left indent_script_item ((1,[]),false,[],[]) cmds + in + let indented_cmds = List.rev (indented_cmds) in + msg_notice (v 0 (Pp.prlist_with_sep Pp.fnl (fun x -> x) indented_cmds)) + with Proof_global.NoCurrentProof -> () + +let () = Vernacentries.show_script := show_script + +(* vim:set foldmethod=marker: *) diff --git a/toplevel/stm.mli b/toplevel/stm.mli new file mode 100644 index 000000000..d4898b30f --- /dev/null +++ b/toplevel/stm.mli @@ -0,0 +1,27 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* Vernacexpr.located_vernac_expr -> unit + +(* Evaluates the tip of the current branch *) +val finish : unit -> unit + +(* Evaluates a particular state id (does not move the current tip) *) +val observe : Stateid.state_id -> unit + +(* Joins the entire document. Implies finish, but also checks proofs *) +val join : unit -> unit + +val get_current_state : unit -> Stateid.state_id +val current_proof_depth : unit -> int +val get_all_proof_names : unit -> Names.identifier list +val get_current_proof_name : unit -> Names.identifier option + +val init : unit -> unit diff --git a/toplevel/toplevel.ml b/toplevel/toplevel.ml index 9a2b8840c..dcf2b0b3d 100644 --- a/toplevel/toplevel.ml +++ b/toplevel/toplevel.ml @@ -213,9 +213,9 @@ let make_prompt () = "n |lem1|lem2|lem3| p < " *) let make_emacs_prompt() = - let statnum = string_of_int (Lib.current_command_label ()) in - let dpth = Pfedit.current_proof_depth() in - let pending = Pfedit.get_all_proof_names() in + let statnum = Stateid.string_of_state_id (Stm.get_current_state ()) in + let dpth = Stm.current_proof_depth() in + let pending = Stm.get_all_proof_names() in let pendingprompt = List.fold_left (fun acc x -> acc ^ (if String.is_empty acc then "" else "|") ^ Names.Id.to_string x) @@ -306,14 +306,24 @@ let do_vernac () = msgerrnl (mt ()); if !print_emacs then msgerr (str (top_buffer.prompt())); resynch_buffer top_buffer; - try ignore (Vernac.eval_expr (read_sentence ())) + try Vernac.eval_expr (read_sentence ()); Stm.finish () with | End_of_input | Errors.Quit -> msgerrnl (mt ()); pp_flush(); raise Errors.Quit | Errors.Drop -> (* Last chance *) if Mltop.is_ocaml_top() then raise Errors.Drop else ppnl (str"Error: There is no ML toplevel." ++ fnl ()) - | any -> ppnl (print_toplevel_error any) + | any -> + Format.set_formatter_out_channel stdout; + ppnl (print_toplevel_error any); + pp_flush (); + match Stateid.get_state_id any with + | Some (valid_id,_) -> + let valid = Stateid.int_of_state_id valid_id in + Vernac.eval_expr (Loc.ghost, + (Vernacexpr.VernacStm (Vernacexpr.Command + (Vernacexpr.VernacBackTo valid)))) + | _ -> () (** Main coq loop : read vernacular expressions until Drop is entered. Ctrl-C is handled internally as Sys.Break instead of aborting Coq. @@ -330,5 +340,7 @@ let rec loop () = | Errors.Drop -> () | Errors.Quit -> exit 0 | any -> - msgerrnl (str"Anomaly: main loop exited. Please report."); - loop () + msgerrnl (str"Anomaly: main loop exited with exception: " ++ + str (Printexc.to_string any) ++ + fnl() ++ str"Please report."); + loop () diff --git a/toplevel/toplevel.mllib b/toplevel/toplevel.mllib index b27c7588d..d58df57f2 100644 --- a/toplevel/toplevel.mllib +++ b/toplevel/toplevel.mllib @@ -2,6 +2,9 @@ Himsg Cerrors Class Locality +Vernacexpr +Ppextra +Ppvernac Metasyntax Auto_ind_decl Search @@ -13,11 +16,12 @@ Command Classes Record Ppvernac -Backtrack Vernacinterp Mltop Vernacentries Whelp +Vernac_classifier +Stm Vernac Ide_slave Toplevel diff --git a/toplevel/vernac.ml b/toplevel/vernac.ml index fcffb343e..e05ced2b8 100644 --- a/toplevel/vernac.ml +++ b/toplevel/vernac.ml @@ -116,53 +116,6 @@ let disable_drop = function let user_error loc s = Errors.user_err_loc (loc,"_",str s) -(** Timeout handling *) - -(** A global default timeout, controled by option "Set Default Timeout n". - Use "Unset Default Timeout" to deactivate it (or set it to 0). *) - -let default_timeout = ref None - -let _ = - Goptions.declare_int_option - { Goptions.optsync = true; - Goptions.optdepr = false; - Goptions.optname = "the default timeout"; - Goptions.optkey = ["Default";"Timeout"]; - Goptions.optread = (fun () -> !default_timeout); - Goptions.optwrite = ((:=) default_timeout) } - -(** When interpreting a command, the current timeout is initially - the default one, but may be modified locally by a Timeout command. *) - -let current_timeout = ref None - -(** Installing and de-installing a timer. - Note: according to ocaml documentation, Unix.alarm isn't available - for native win32. *) - -let timeout_handler _ = raise Timeout - -let set_timeout n = - let psh = - Sys.signal Sys.sigalrm (Sys.Signal_handle timeout_handler) in - ignore (Unix.alarm n); - Some psh - -let default_set_timeout () = - match !current_timeout with - | Some n -> set_timeout n - | None -> None - -let restore_timeout = function - | None -> () - | Some psh -> - (* stop alarm *) - ignore(Unix.alarm 0); - (* restore handler *) - Sys.set_signal Sys.sigalrm psh - - (* Open an utf-8 encoded file and skip the byte-order mark if any *) let open_utf8_file_in fname = @@ -262,7 +215,30 @@ let restore_translator_coqdoc (ch,cl,cs,coqdocstate) = Lexer.restore_com_state cs; Lexer.restore_location_table coqdocstate -let rec vernac_com interpfun checknav (loc,com) = +(* For coqtop -time, we display the position in the file, + and a glimpse of the executed command *) + +let display_cmd_header loc com = + let shorten s = try (String.sub s 0 30)^"..." with _ -> s in + let noblank s = + for i = 0 to String.length s - 1 do + match s.[i] with + | ' ' | '\n' | '\t' | '\r' -> s.[i] <- '~' + | _ -> () + done; + s + in + let (start,stop) = Loc.unloc loc in + let safe_pr_vernac x = + try Ppvernac.pr_vernac x + with e -> str (Printexc.to_string e) in + let cmd = noblank (shorten (string_of_ppcmds (safe_pr_vernac com))) + in + Pp.pp (str "Chars " ++ int start ++ str " - " ++ int stop ++ + str (" ["^cmd^"] ")); + Pp.flush_all () + +let rec vernac_com verbosely checknav (loc,com) = let rec interp = function | VernacLoad (verbosely, fname) -> let fname = Envars.expand_path_macros ~warn:(fun x -> msg_warning (str x)) fname in @@ -286,39 +262,10 @@ let rec vernac_com interpfun checknav (loc,com) = raise reraise end - | VernacList l -> - List.iter (fun (_,v) -> interp v) l - | v when !just_parsing -> () - | VernacFail v -> - begin - try - (* If the command actually works, ignore its effects on the state *) - States.with_state_protection - (fun v -> ignore (interp v); raise HasNotFailed) v - with - | HasNotFailed -> - errorlabstrm "Fail" (str "The command has not failed !") - | e when Errors.noncritical e -> (* In particular e is no anomaly *) - if_verbose msg_info - (str "The command has indeed failed with message:" ++ - fnl () ++ str "=> " ++ hov 0 (Errors.print e)) - end - - | VernacTime v -> - let tstart = System.get_time() in - interp v; - let tend = System.get_time() in - let msg = if !time then "" else "Finished transaction in " in - msg_info (str msg ++ System.fmt_time_difference tstart tend) - - | VernacTimeout(n,v) -> - current_timeout := Some n; - interp v - - | v -> - let rollback = + | v -> Stm.process_transaction verbosely (loc,v) + (* FIXME elsewhere let rollback = if !Flags.batch_mode then States.without_rollback else States.with_heavy_rollback in @@ -330,12 +277,12 @@ let rec vernac_com interpfun checknav (loc,com) = let reraise = Errors.push reraise in restore_timeout psh; raise reraise + *) in try checknav loc com; - current_timeout := !default_timeout; if do_beautify () then pr_new_syntax loc (Some com); - if !time then display_cmd_header loc com; + if !Flags.time then display_cmd_header loc com; let com = if !time then VernacTime com else com in interp com with reraise -> @@ -347,20 +294,10 @@ let rec vernac_com interpfun checknav (loc,com) = and read_vernac_file verbosely s = Flags.make_warn verbosely; - let interpfun = - if verbosely then Vernacentries.interp - else Flags.silently Vernacentries.interp - in let checknav loc cmd = if is_navigation_vernac cmd && not (is_reset cmd) then user_error loc "Navigation commands forbidden in files" in - let end_inner_command cmd = - if !atomic_load || is_reset cmd then - Lib.mark_end_of_command () (* for Reset in coqc or coqtop -l *) - else - Backtrack.mark_command cmd; (* for Show Script, cf bug #2820 *) - in let (in_chan, fname, input) = open_file_twice_if verbosely s in try @@ -368,8 +305,7 @@ and read_vernac_file verbosely s = * raised, which means that we raised the end of the file being loaded *) while true do let loc_ast = parse_sentence input in - vernac_com interpfun checknav loc_ast; - end_inner_command (snd loc_ast); + vernac_com verbosely checknav loc_ast; pp_flush () done with any -> (* whatever the exception *) @@ -378,6 +314,7 @@ and read_vernac_file verbosely s = close_input in_chan input; (* we must close the file first *) match e with | End_of_input -> + Stm.join (); if do_beautify () then pr_new_syntax (Loc.make_loc (max_int,max_int)) None | _ -> raise_with_file fname (disable_drop e) @@ -393,10 +330,7 @@ let checknav loc ast = if is_deep_navigation_vernac ast then user_error loc "Navigation commands forbidden in nested commands" -let eval_expr ?(preserving=false) loc_ast = - vernac_com Vernacentries.interp checknav loc_ast; - if not preserving && not (is_navigation_vernac (snd loc_ast)) then - Backtrack.mark_command (snd loc_ast) +let eval_expr loc_ast = vernac_com true checknav loc_ast (* XML output hooks *) let (f_xml_start_library, xml_start_library) = Hook.make ~default:ignore () @@ -407,7 +341,6 @@ let load_vernac verb file = chan_beautify := if !Flags.beautify_file then open_out (file^beautify_suffix) else stdout; try - Lib.mark_end_of_command (); (* in case we're still in coqtop init *) read_vernac_file verb file; if !Flags.beautify_file then close_out !chan_beautify; with any -> @@ -428,3 +361,4 @@ let compile verbosely f = if !Flags.xml_export then Hook.get f_xml_end_library (); Library.save_library_to ldir (long_f_dot_v ^ "o"); Dumpglob.end_dump_glob () + diff --git a/toplevel/vernac.mli b/toplevel/vernac.mli index bda42dfa6..d18571990 100644 --- a/toplevel/vernac.mli +++ b/toplevel/vernac.mli @@ -21,12 +21,7 @@ exception End_of_input val just_parsing : bool ref -(** [eval_expr] executes one vernacular command. By default the command is - considered as non-state-preserving, in which case we add it to the - Backtrack stack (triggering a save of a frozen state and the generation - of a new state label). An example of state-preserving command is one coming - from the query panel of Coqide. *) -val eval_expr : ?preserving:bool -> Loc.t * Vernacexpr.vernac_expr -> unit +val eval_expr : Loc.t * Vernacexpr.vernac_expr -> unit (** Set XML hooks *) val xml_start_library : (unit -> unit) Hook.t @@ -42,7 +37,6 @@ val load_vernac : bool -> string -> unit val compile : bool -> string -> unit - val is_navigation_vernac : Vernacexpr.vernac_expr -> bool (** Should we display timings for each command ? *) diff --git a/toplevel/vernac_classifier.ml b/toplevel/vernac_classifier.ml new file mode 100644 index 000000000..d7b9553a0 --- /dev/null +++ b/toplevel/vernac_classifier.ml @@ -0,0 +1,164 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* "Unknown" + | VtStartProof _ -> "StartProof" + | VtSideff -> "Sideff" + | VtQed _ -> "Qed" + | VtProofStep -> "ProofStep" + | VtQuery b -> "Query" ^ string_of_in_script b + | VtStm ((VtFinish|VtJoinDocument|VtObserve _), b) -> + "Stm" ^ string_of_in_script b + | VtStm (VtBack _, b) -> "Stm Back" ^ string_of_in_script b + +let string_of_vernac_when = function + | VtLater -> "Later" + | VtNow -> "Now" + +let string_of_vernac_classification (t,w) = + string_of_vernac_type t ^ " " ^ string_of_vernac_when w + +(* Since the set of vernaculars is extensible, also the classification function + has to be. *) +let classifiers = Summary.ref ~name:"vernac_classifier" [] +let declare_vernac_classifier s (f : vernac_expr -> vernac_classification) = + classifiers := !classifiers @ [s,f] + +let elide_part_of_script_and_now (a, _) = + match a with + | VtQuery _ -> VtQuery false, VtNow + | VtStm (x, _) -> VtStm (x, false), VtNow + | x -> x, VtNow + +let rec classify_vernac e = + let static_classifier e = match e with + (* Stm *) + | VernacStm Finish -> VtStm (VtFinish, true), VtNow + | VernacStm JoinDocument -> VtStm (VtJoinDocument, true), VtNow + | VernacStm (Observe id) -> VtStm (VtObserve id, true), VtNow + | VernacStm (Command x) -> elide_part_of_script_and_now (classify_vernac x) + (* Impossible, Vernac handles these *) + | VernacList _ -> anomaly (str "VernacList not handled by Vernac") + | VernacLoad _ -> anomaly (str "VernacLoad not handled by Vernac") + (* Nested vernac exprs *) + | VernacProgram e -> classify_vernac e + | VernacLocal (_,e) -> classify_vernac e + | VernacTimeout (_,e) -> classify_vernac e + | VernacTime e -> classify_vernac e + | VernacFail e -> (* Fail Qed or Fail Lemma must not join/fork the DAG *) + (match classify_vernac e with + | (VtQuery _ | VtProofStep _ | VtSideff _ | VtStm _), _ as x -> x + | VtQed _, _ -> VtProofStep, VtNow + | (VtStartProof _ | VtUnknown), _ -> VtUnknown, VtNow) + (* Qed *) + | VernacEndProof Admitted | VernacAbort _ -> VtQed DropProof, VtLater + | VernacEndProof _ | VernacExactProof _ -> VtQed KeepProof, VtLater + (* Query *) + | VernacShow _ | VernacPrint _ | VernacSearch _ | VernacLocate _ + | VernacCheckMayEval _ -> VtQuery true, VtLater + (* ProofStep *) + | VernacProof _ + | VernacBullet _ + | VernacFocus _ | VernacUnfocus _ + | VernacSubproof _ | VernacEndSubproof _ + | VernacSolve _ + | VernacCheckGuard _ + | VernacUnfocused _ + | VernacSolveExistential _ -> VtProofStep, VtLater + (* StartProof *) + | VernacDefinition (_,(_,i),ProveBody _) -> + VtStartProof("Classic",[i]), VtLater + | VernacStartTheoremProof (_,l,_) -> + let names = + CList.map_filter (function (Some(_,i), _) -> Some i | _ -> None) l in + VtStartProof ("Classic", names), VtLater + | VernacGoal _ -> VtStartProof ("Classic",[]), VtLater + | VernacFixpoint (_,l) + when List.exists (fun ((_,_,_,_,p),_) -> p = None) l -> + VtStartProof ("Classic", + List.map (fun (((_,id),_,_,_,_),_) -> id) l), VtLater + | VernacCoFixpoint (_,l) + when List.exists (fun ((_,_,_,p),_) -> p = None) l -> + VtStartProof ("Classic", + List.map (fun (((_,id),_,_,_),_) -> id) l), VtLater + (* Sideff: apply to all open branches. usually run on master only *) + | VernacAssumption _ + | VernacDefinition (_,_,DefineBody _) + | VernacFixpoint _ | VernacCoFixpoint _ + | VernacInductive _ + | VernacScheme _ | VernacCombinedScheme _ + | VernacBeginSection _ + | VernacCanonical _ | VernacCoercion _ | VernacIdentityCoercion _ + | VernacAddLoadPath _ | VernacRemoveLoadPath _ | VernacAddMLPath _ + | VernacChdir _ + | VernacCreateHintDb _ | VernacRemoveHints _ | VernacHints _ + | VernacDeclareImplicits _ | VernacArguments _ | VernacArgumentsScope _ + | VernacReserve _ + | VernacGeneralizable _ + | VernacSetOpacity _ | VernacSetStrategy _ + | VernacUnsetOption _ | VernacSetOption _ + | VernacAddOption _ | VernacRemoveOption _ + | VernacMemOption _ | VernacPrintOption _ + | VernacGlobalCheck _ + | VernacDeclareReduction _ + | VernacDeclareClass _ | VernacDeclareInstances _ + | VernacComments _ -> VtSideff, VtLater + (* (Local) Notations have to disappear *) + | VernacEndSegment _ -> VtSideff, VtNow + (* Modules with parameters have to be executed: can import notations *) + | VernacDeclareModule (_,_,bl,_) + | VernacDefineModule (_,_,bl,_,_) + | VernacDeclareModuleType (_,bl,_,_) -> + VtSideff, if bl = [] then VtLater else VtNow + (* These commands alter the parser *) + | VernacOpenCloseScope _ | VernacDelimiters _ | VernacBindScope _ + | VernacInfix _ | VernacNotation _ | VernacSyntaxExtension _ + | VernacSyntacticDefinition _ + | VernacDeclareTacticDefinition _ | VernacTacticNotation _ + | VernacRequire _ | VernacImport _ | VernacInclude _ + | VernacDeclareMLModule _ + | VernacContext _ (* TASSI: unsure *) + | VernacProofMode _ + | VernacRequireFrom _ -> VtSideff, VtNow + (* These are ambiguous *) + | VernacInstance _ -> VtUnknown, VtNow + (* Stm will install a new classifier to handle these *) + | VernacBack _ | VernacAbortAll + | VernacUndoTo _ | VernacUndo _ + | VernacResetName _ | VernacResetInitial _ + | VernacBacktrack _ | VernacBackTo _ | VernacRestart -> !undo_classifier e + (* What are these? *) + | VernacNop + | VernacToplevelControl _ + | VernacRestoreState _ + | VernacWriteState _ -> VtUnknown, VtNow + (* Plugins should classify their commands *) + | VernacExtend _ -> VtUnknown, VtNow in + let rec extended_classifier = function + | [] -> static_classifier e + | (name,f)::fs -> + try + match f e with + | VtUnknown, _ -> extended_classifier fs + | x -> x + with e when Errors.noncritical e -> + let e = Errors.push e in + msg_warning(str"Exception raised by classifier: " ++ str name); + raise e + + in + extended_classifier !classifiers + diff --git a/toplevel/vernac_classifier.mli b/toplevel/vernac_classifier.mli new file mode 100644 index 000000000..06535d512 --- /dev/null +++ b/toplevel/vernac_classifier.mli @@ -0,0 +1,20 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* string + +(** What does a vernacular do *) +val classify_vernac : vernac_expr -> vernac_classification + +(** Install an additional vernacular classifier (that has precedence over + all classifiers already installed) *) +val declare_vernac_classifier : + string -> (vernac_expr -> vernac_classification) -> unit diff --git a/toplevel/vernacentries.ml b/toplevel/vernacentries.ml index 88da26e83..706b757db 100644 --- a/toplevel/vernacentries.ml +++ b/toplevel/vernacentries.ml @@ -34,6 +34,10 @@ open Declaremods open Misctypes open Locality +let debug = false +let prerr_endline = + if debug then prerr_endline else fun _ -> () + (* Misc *) let cl_of_qualid = function @@ -58,60 +62,6 @@ let show_node () = could, possibly, be cleaned away. (Feb. 2010) *) () -(* indentation code for Show Script, initially contributed - by D. de Rauglaudre *) - -let indent_script_item ((ng1,ngl1),nl,beginend,ppl) (cmd,ng) = - (* ng1 : number of goals remaining at the current level (before cmd) - ngl1 : stack of previous levels with their remaining goals - ng : number of goals after the execution of cmd - beginend : special indentation stack for { } *) - let ngprev = List.fold_left (+) ng1 ngl1 in - let new_ngl = - if ng > ngprev then - (* We've branched *) - (ng - ngprev + 1, ng1 - 1 :: ngl1) - else if ng < ngprev then - (* A subgoal have been solved. Let's compute the new current level - by discarding all levels with 0 remaining goals. *) - let _ = assert (Int.equal ng (ngprev - 1)) in - let rec loop = function - | (0, ng2::ngl2) -> loop (ng2,ngl2) - | p -> p - in loop (ng1-1, ngl1) - else - (* Standard case, same goal number as before *) - (ng1, ngl1) - in - (* When a subgoal have been solved, separate this block by an empty line *) - let new_nl = (ng < ngprev) - in - (* Indentation depth *) - let ind = List.length ngl1 - in - (* Some special handling of bullets and { }, to get a nicer display *) - let pred n = max 0 (n-1) in - let ind, nl, new_beginend = match cmd with - | VernacSubproof _ -> pred ind, nl, (pred ind)::beginend - | VernacEndSubproof -> List.hd beginend, false, List.tl beginend - | VernacBullet _ -> pred ind, nl, beginend - | _ -> ind, nl, beginend - in - let pp = - (if nl then fnl () else mt ()) ++ - (hov (ind+1) (str (String.make ind ' ') ++ Ppvernac.pr_vernac cmd)) - in - (new_ngl, new_nl, new_beginend, pp :: ppl) - -let show_script () = - let prf = Pfedit.get_current_proof_name () in - let cmds = Backtrack.get_script prf in - let _,_,_,indented_cmds = - List.fold_left indent_script_item ((1,[]),false,[],[]) cmds - in - let indented_cmds = List.rev (indented_cmds) in - msg_notice (v 0 (Pp.prlist_with_sep Pp.fnl (fun x -> x) indented_cmds)) - let show_thesis () = msg_error (anomaly (Pp.str "TODO") ) @@ -130,7 +80,9 @@ let enable_goal_printing = ref true let print_subgoals () = if !enable_goal_printing && is_verbose () - then msg_notice (pr_open_subgoals ()) + then begin + msg_notice (pr_open_subgoals ()) + end let try_print_subgoals () = Pp.flush_all(); @@ -141,8 +93,8 @@ let try_print_subgoals () = let show_intro all = let pf = get_pftreestate() in - let {Evd.it=gls ; sigma=sigma} = Proof.V82.subgoals pf in - let gl = {Evd.it=List.hd gls ; sigma = sigma} in + let {Evd.it=gls ; sigma=sigma; eff=eff} = Proof.V82.subgoals pf in + let gl = {Evd.it=List.hd gls ; sigma = sigma; eff=eff} in let l,_= decompose_prod_assum (strip_outer_cast (pf_concl gl)) in if all then @@ -502,21 +454,20 @@ let vernac_start_proof kind l lettop = start_proof_and_print (Global, Proof kind) l no_hook let qed_display_script = ref true +let show_script = ref (fun ?proof () -> ()) -let vernac_end_proof = function +let vernac_end_proof ?proof = function | Admitted -> - Backtrack.mark_unreachable [Pfedit.get_current_proof_name ()]; admit (); Pp.feedback Interface.AddedAxiom | Proved (is_opaque,idopt) -> - let prf = Pfedit.get_current_proof_name () in - if is_verbose () && !qed_display_script then show_script (); + if is_verbose () && !qed_display_script then !show_script ?proof (); begin match idopt with - | None -> save_named is_opaque - | Some ((_,id),None) -> save_anonymous is_opaque id - | Some ((_,id),Some kind) -> save_anonymous_with_strength kind is_opaque id - end; - Backtrack.mark_unreachable [prf] + | None -> save_named ?proof is_opaque + | Some ((_,id),None) -> save_anonymous ?proof is_opaque id + | Some ((_,id),Some kind) -> + save_anonymous_with_strength ?proof kind is_opaque id + end (* A stupid macro that should be replaced by ``Exact c. Save.'' all along the theories [??] *) @@ -524,10 +475,8 @@ let vernac_end_proof = function let vernac_exact_proof c = (* spiwack: for simplicity I do not enforce that "Proof proof_term" is called only at the begining of a proof. *) - let prf = Pfedit.get_current_proof_name () in by (Tactics.exact_proof c); - save_named true; - Backtrack.mark_unreachable [prf] + save_named true let vernac_assumption locality (local, kind) l nl = let local = enforce_locality_exp locality local in @@ -800,9 +749,10 @@ let vernac_identity_coercion locality local id qids qidt = (* Type classes *) let vernac_instance abst locality sup inst props pri = - let glob = not (make_section_locality locality) in + let global = not (make_section_locality locality) in Dumpglob.dump_constraint inst false "inst"; - ignore(Classes.new_instance ~abstract:abst ~global:glob sup inst props pri) + ignore(Classes.new_instance + ~abstract:abst ~global sup inst props pri) let vernac_context l = if not (Classes.context l) then Pp.feedback Interface.AddedAxiom @@ -824,15 +774,15 @@ let focus_command_cond = Proof.no_cond command_focus let vernac_solve n tcom b = if not (refining ()) then error "Unknown command of the non proof-editing mode."; - let p = Proof_global.give_me_the_proof () in - Proof.transaction p begin fun () -> - solve_nth n (Tacinterp.hide_interp tcom None) ~with_end_tac:b; + Proof_global.with_current_proof (fun etac p -> + let with_end_tac = if b then Some etac else None in + let p = solve_nth n (Tacinterp.hide_interp tcom None) ?with_end_tac p in (* in case a strict subtree was completed, go back to the top of the prooftree *) - Proof_global.maximal_unfocus command_focus p; + let p = Proof.maximal_unfocus command_focus p in + p); print_subgoals() - end - +;; (* A command which should be a tactic. It has been added by Christine to patch an error in the design of the proof @@ -1532,126 +1482,23 @@ let vernac_register id r = match r with | RegisterInline -> Global.register_inline kn -(****************) -(* Backtracking *) - -(** NB: these commands are now forbidden in non-interactive use, - e.g. inside VernacLoad, VernacList, ... *) - -let vernac_backto lbl = - try - let lbl' = Backtrack.backto lbl in - if not (Int.equal lbl lbl') then - Pp.msg_warning - (str "Actually back to state "++ Pp.int lbl' ++ str "."); - try_print_subgoals () - with Backtrack.Invalid -> error "Invalid backtrack." - -let vernac_back n = - try - let extra = Backtrack.back n in - if not (Int.equal extra 0) then - Pp.msg_warning - (str "Actually back by " ++ Pp.int (extra+n) ++ str " steps."); - try_print_subgoals () - with Backtrack.Invalid -> error "Invalid backtrack." - -let vernac_reset_name id = - try - let globalized = - try - let gr = Smartlocate.global_with_alias (Ident id) in - Dumpglob.add_glob (fst id) gr; - true - with e when Errors.noncritical e -> false in - - if not globalized then begin - try begin match Lib.find_opening_node (snd id) with - | Lib.OpenedSection _ -> Dumpglob.dump_reference (fst id) - (DirPath.to_string (Lib.current_dirpath true)) "<>" "sec"; - (* Might be nice to implement module cases, too.... *) - | _ -> () - end with UserError _ -> () - end; - - if Backtrack.is_active () then - (Backtrack.reset_name id; try_print_subgoals ()) - else - (* When compiling files, Reset is now allowed again - as asked by A. Chlipala. we emulate a simple reset, - that discards all proofs. *) - let lbl = Lib.label_before_name id in - Pfedit.delete_all_proofs (); - Pp.msg_warning (str "Reset command occurred in non-interactive mode."); - Lib.reset_label lbl - with Backtrack.Invalid | Not_found -> error "Invalid Reset." - - -let vernac_reset_initial () = - if Backtrack.is_active () then - Backtrack.reset_initial () - else begin - Pp.msg_warning (str "Reset command occurred in non-interactive mode."); - Lib.raw_reset_initial () - end - -(* For compatibility with ProofGeneral: *) - -let vernac_backtrack snum pnum naborts = - Backtrack.backtrack snum pnum naborts; - try_print_subgoals () - - (********************) (* Proof management *) -let vernac_abort = function - | None -> - Backtrack.mark_unreachable [Pfedit.get_current_proof_name ()]; - delete_current_proof (); - if_verbose msg_info (str "Current goal aborted") - | Some id -> - Backtrack.mark_unreachable [snd id]; - delete_proof id; - let s = Id.to_string (snd id) in - if_verbose msg_info (str ("Goal "^s^" aborted")) - -let vernac_abort_all () = - if refining() then begin - Backtrack.mark_unreachable (Pfedit.get_all_proof_names ()); - delete_all_proofs (); - msg_info (str "Current goals aborted") - end else - error "No proof-editing in progress." - -let vernac_restart () = - Backtrack.mark_unreachable [Pfedit.get_current_proof_name ()]; - restart_proof(); print_subgoals () - -let vernac_undo n = - let d = Pfedit.current_proof_depth () - n in - Backtrack.mark_unreachable ~after:d [Pfedit.get_current_proof_name ()]; - Pfedit.undo n; print_subgoals () - -let vernac_undoto n = - Backtrack.mark_unreachable ~after:n [Pfedit.get_current_proof_name ()]; - Pfedit.undo_todepth n; - print_subgoals () - let vernac_focus gln = - let p = Proof_global.give_me_the_proof () in - let n = match gln with None -> 1 | Some n -> n in - if Int.equal n 0 then - Errors.error "Invalid goal number: 0. Goal numbering starts with 1." - else - try Proof.focus focus_command_cond () n p; print_subgoals () - with Proofview.IndexOutOfRange -> - Errors.error "No such unproven subgoal." + Proof_global.with_current_proof (fun _ p -> + match gln with + | None -> Proof.focus focus_command_cond () 1 p + | Some 0 -> + Errors.error "Invalid goal number: 0. Goal numbering starts with 1." + | Some n -> + Proof.focus focus_command_cond () n p); + print_subgoals () (* Unfocuses one step in the focus stack. *) let vernac_unfocus () = - let p = Proof_global.give_me_the_proof () in - Proof.unfocus command_focus p; print_subgoals () + Proof_global.with_current_proof (fun _ p -> Proof.unfocus command_focus p ()); + print_subgoals () (* Checks that a proof is fully unfocused. Raises an error if not. *) let vernac_unfocused () = @@ -1672,22 +1519,20 @@ let subproof_kind = Proof.new_focus_kind () let subproof_cond = Proof.done_cond subproof_kind let vernac_subproof gln = - let p = Proof_global.give_me_the_proof () in - begin match gln with - | None -> Proof.focus subproof_cond () 1 p - | Some n -> Proof.focus subproof_cond () n p - end ; + Proof_global.with_current_proof (fun _ p -> + match gln with + | None -> Proof.focus subproof_cond () 1 p + | Some n -> Proof.focus subproof_cond () n p); print_subgoals () let vernac_end_subproof () = - let p = Proof_global.give_me_the_proof () in - Proof.unfocus subproof_kind p ; print_subgoals () + Proof_global.with_current_proof (fun _ p -> Proof.unfocus subproof_kind p ()); + print_subgoals () let vernac_bullet (bullet:Proof_global.Bullet.t) = - let p = Proof_global.give_me_the_proof () in - Proof.transaction p - (fun () -> Proof_global.Bullet.put p bullet); + Proof_global.with_current_proof (fun _ p -> + Proof_global.Bullet.put p bullet); (* Makes the focus visible in emacs by re-printing the goal. *) if !Flags.print_emacs then print_subgoals () @@ -1706,7 +1551,7 @@ let vernac_show = function Constrextern.with_implicits msg_notice (pr_nth_open_subgoal n) | ShowProof -> show_proof () | ShowNode -> show_node () - | ShowScript -> show_script () + | ShowScript -> !show_script () | ShowExistentials -> show_top_evars () | ShowTree -> show_prooftree () | ShowProofNames -> @@ -1733,10 +1578,17 @@ let vernac_check_guard () = (* "locality" is the prefix "Local" attribute, while the "local" component * is the outdated/deprecated "Local" attribute of some vernacular commands * still parsed as the obsolete_locality grammar entry for retrocompatibility *) -let interp locality c = match c with - (* Control (done in vernac) *) - | (VernacTime _|VernacList _|VernacLoad _|VernacTimeout _|VernacFail _) -> - assert false +let interp ?proof locality c = + prerr_endline ("interpreting: " ^ Pp.string_of_ppcmds (Ppvernac.pr_vernac c)); + match c with + (* Done in vernac *) + | (VernacList _|VernacLoad _) -> assert false + + (* Done later in this file *) + | VernacFail _ -> assert false + | VernacTime _ -> assert false + | VernacTimeout _ -> assert false + | VernacStm _ -> assert false (* Syntax *) | VernacTacticNotation (n,r,e) -> @@ -1754,7 +1606,7 @@ let interp locality c = match c with (* Gallina *) | VernacDefinition (k,lid,d) -> vernac_definition locality k lid d | VernacStartTheoremProof (k,l,top) -> vernac_start_proof k l top - | VernacEndProof e -> vernac_end_proof e + | VernacEndProof e -> vernac_end_proof ?proof e | VernacExactProof c -> vernac_exact_proof c | VernacAssumption (stre,nl,l) -> vernac_assumption locality stre l nl | VernacInductive (finite,infer,l) -> vernac_inductive finite infer l @@ -1808,10 +1660,10 @@ let interp locality c = match c with | VernacRestoreState s -> vernac_restore_state s (* Resetting *) - | VernacResetName id -> vernac_reset_name id - | VernacResetInitial -> vernac_reset_initial () - | VernacBack n -> vernac_back n - | VernacBackTo n -> vernac_backto n + | VernacResetName _ -> anomaly (str "VernacResetName not handled by Stm") + | VernacResetInitial -> anomaly (str "VernacResetInitial not handled by Stm") + | VernacBack _ -> anomaly (str "VernacBack not handled by Stm") + | VernacBackTo _ -> anomaly (str "VernacBackTo not handled by Stm") (* Commands *) | VernacDeclareTacticDefinition def -> @@ -1848,12 +1700,12 @@ let interp locality c = match c with (* Proof management *) | VernacGoal t -> vernac_start_proof Theorem [None,([],t,None)] false - | VernacAbort id -> vernac_abort id - | VernacAbortAll -> vernac_abort_all () - | VernacRestart -> vernac_restart () - | VernacUndo n -> vernac_undo n - | VernacUndoTo n -> vernac_undoto n - | VernacBacktrack (snum,pnum,naborts) -> vernac_backtrack snum pnum naborts + | VernacAbort id -> anomaly (str "VernacAbort not handled by Stm") + | VernacAbortAll -> anomaly (str "VernacAbortAll not handled by Stm") + | VernacRestart -> anomaly (str "VernacRestart not handled by Stm") + | VernacUndo _ -> anomaly (str "VernacUndo not handled by Stm") + | VernacUndoTo _ -> anomaly (str "VernacUndoTo not handled by Stm") + | VernacBacktrack _ -> anomaly (str "VernacBacktrack not handled by Stm") | VernacFocus n -> vernac_focus n | VernacUnfocus -> vernac_unfocus () | VernacUnfocused -> vernac_unfocused () @@ -1902,24 +1754,119 @@ let check_vernac_supports_locality c l = | VernacExtend _ ) -> () | Some _, _ -> Errors.error "This command does not support Locality" -let interp c = +(** A global default timeout, controled by option "Set Default Timeout n". + Use "Unset Default Timeout" to deactivate it (or set it to 0). *) + +let default_timeout = ref None + +let _ = + Goptions.declare_int_option + { Goptions.optsync = true; + Goptions.optdepr = false; + Goptions.optname = "the default timeout"; + Goptions.optkey = ["Default";"Timeout"]; + Goptions.optread = (fun () -> !default_timeout); + Goptions.optwrite = ((:=) default_timeout) } + +(** When interpreting a command, the current timeout is initially + the default one, but may be modified locally by a Timeout command. *) + +let current_timeout = ref None + +(** Installing and de-installing a timer. + Note: according to ocaml documentation, Unix.alarm isn't available + for native win32. *) + +let timeout_handler _ = raise Timeout + +let set_timeout n = + let psh = + Sys.signal Sys.sigalrm (Sys.Signal_handle timeout_handler) in + ignore (Unix.alarm n); + Some psh + +let default_set_timeout () = + match !current_timeout with + | Some n -> set_timeout n + | None -> None + +let restore_timeout = function + | None -> () + | Some psh -> + (* stop alarm *) + ignore(Unix.alarm 0); + (* restore handler *) + Sys.set_signal Sys.sigalrm psh + +let locate_if_not_already loc exn = + match Loc.get_loc exn with + | None -> Loc.add_loc exn loc + | Some l -> if Loc.is_ghost l then Loc.add_loc exn loc else exn + +exception HasNotFailed +exception HasFailed of string + +let interp ?(verbosely=true) ?proof (loc,c) = let orig_program_mode = Flags.is_program_mode () in let rec aux ?locality isprogcmd = function | VernacProgram c when not isprogcmd -> aux ?locality true c | VernacProgram _ -> Errors.error "Program mode specified twice" | VernacLocal (b, c) when locality = None -> aux ~locality:b isprogcmd c | VernacLocal _ -> Errors.error "Locality specified twice" + | VernacStm (Command c) -> aux ?locality isprogcmd c + | VernacStm _ -> assert false (* Done by Stm *) + | VernacFail v -> + begin try + (* If the command actually works, ignore its effects on the state. + * Note that error has to be printed in the right state, hence + * within the purified function *) + Future.purify + (fun v -> + try + aux ?locality isprogcmd v; + raise HasNotFailed + with + | HasNotFailed as e -> raise e + | e -> raise (HasFailed (Pp.string_of_ppcmds + (Errors.print (Cerrors.process_vernac_interp_error e))))) + v + with e when Errors.noncritical e -> + let e = Errors.push e in + match e with + | HasNotFailed -> + errorlabstrm "Fail" (str "The command has not failed!") + | HasFailed msg -> + if_verbose msgnl + (str "The command has indeed failed with message:" ++ + fnl () ++ str "=> " ++ hov 0 (str msg)) + | _ -> assert false + end + | VernacTimeout (n,v) -> + current_timeout := Some n; + aux ?locality isprogcmd v + | VernacTime v -> + let tstart = System.get_time() in + aux ?locality isprogcmd v; + let tend = System.get_time() in + let msg = if !Flags.time then "" else "Finished transaction in " in + msg_info (str msg ++ System.fmt_time_difference tstart tend) | c -> check_vernac_supports_locality c locality; Obligations.set_program_mode isprogcmd; + let psh = default_set_timeout () in try - interp locality c; + if verbosely then Flags.verbosely (interp ?proof locality) c + else Flags.silently (interp ?proof locality) c; + restore_timeout psh; if orig_program_mode || not !Flags.program_mode || isprogcmd then Flags.program_mode := orig_program_mode with - | reraise -> + | reraise when Errors.noncritical reraise -> let e = Errors.push reraise in + let e = locate_if_not_already loc e in + restore_timeout psh; Flags.program_mode := orig_program_mode; raise e in aux false c + diff --git a/toplevel/vernacentries.mli b/toplevel/vernacentries.mli index 5a6550d56..3f6d45a04 100644 --- a/toplevel/vernacentries.mli +++ b/toplevel/vernacentries.mli @@ -17,7 +17,7 @@ val dump_global : Libnames.reference or_by_notation -> unit (** Vernacular entries *) -val show_script : unit -> unit +val show_script : (?proof:Proof_global.closed_proof -> unit -> unit) ref val show_prooftree : unit -> unit val show_node : unit -> unit @@ -27,12 +27,16 @@ val show_node : unit -> unit val get_current_context_of_args : int option -> Evd.evar_map * Environ.env (** The main interpretation function of vernacular expressions *) -val interp : Vernacexpr.vernac_expr -> unit +val interp : + ?verbosely:bool -> + ?proof:Proof_global.closed_proof -> + Loc.t * Vernacexpr.vernac_expr -> unit (** Print subgoals when the verbose flag is on. Meant to be used inside vernac commands from plugins. *) val print_subgoals : unit -> unit +val try_print_subgoals : unit -> unit (** The printing of goals via [print_subgoals] or during [interp] can be controlled by the following flag. @@ -53,3 +57,6 @@ val qed_display_script : bool ref a known inductive type. *) val make_cases : string -> string list list + +val vernac_end_proof : + ?proof:Proof_global.closed_proof -> Vernacexpr.proof_end -> unit diff --git a/toplevel/whelp.ml4 b/toplevel/whelp.ml4 index 8929fb32c..8185e5702 100644 --- a/toplevel/whelp.ml4 +++ b/toplevel/whelp.ml4 @@ -192,8 +192,8 @@ let whelp_elim ind = send_whelp "elim" (make_string uri_of_global (IndRef ind)) let on_goal f = - let { Evd.it=goals ; sigma=sigma } = Proof.V82.subgoals (get_pftreestate ()) in - let gls = { Evd.it=List.hd goals ; sigma = sigma } in + let gls = Proof.V82.subgoals (get_pftreestate ()) in + let gls = { gls with Evd.it = List.hd gls.Evd.it } in f (Termops.it_mkNamedProd_or_LetIn (pf_concl gls) (pf_hyps gls)) type whelp_request = -- cgit v1.2.3