diff options
-rw-r--r-- | API/API.mli | 8 | ||||
-rw-r--r-- | dev/doc/xml-protocol.md | 5 | ||||
-rw-r--r-- | ide/ide_slave.ml | 1 | ||||
-rw-r--r-- | ide/interface.mli | 5 | ||||
-rw-r--r-- | ide/xmlprotocol.ml | 26 | ||||
-rw-r--r-- | ide/xmlprotocol.mli | 2 | ||||
-rw-r--r-- | intf/vernacexpr.ml | 15 | ||||
-rw-r--r-- | parsing/g_vernac.ml4 | 4 | ||||
-rw-r--r-- | printing/ppvernac.ml | 6 | ||||
-rw-r--r-- | stm/stm.ml | 34 | ||||
-rw-r--r-- | stm/stm.mli | 3 | ||||
-rw-r--r-- | stm/vernac_classifier.ml | 8 | ||||
-rw-r--r-- | tools/fake_ide.ml | 6 | ||||
-rw-r--r-- | toplevel/vernac.ml | 2 | ||||
-rw-r--r-- | vernac/vernacentries.ml | 5 | ||||
-rw-r--r-- | vernac/vernacprop.ml | 3 |
16 files changed, 57 insertions, 76 deletions
diff --git a/API/API.mli b/API/API.mli index 901ed3528..50fbee529 100644 --- a/API/API.mli +++ b/API/API.mli @@ -3676,7 +3676,7 @@ sig | VtProofStep of proof_step | VtProofMode of string | VtQuery of vernac_part_of_script * Feedback.route_id - | VtStm of vernac_control * vernac_part_of_script + | VtBack of vernac_part_of_script * Stateid.t | VtUnknown and vernac_qed_type = | VtKeep @@ -3685,10 +3685,6 @@ sig and vernac_start = string * opacity_guarantee * Names.Id.t list and vernac_sideff_type = Names.Id.t list and vernac_part_of_script = bool - and vernac_control = - | VtWait - | VtJoinDocument - | VtBack of Stateid.t and opacity_guarantee = | GuaranteesOpacity | Doesn'tGuaranteeOpacity @@ -3770,7 +3766,6 @@ sig type option_value type showable type bullet - type stm_vernac type comment type register_kind type locatable @@ -3922,7 +3917,6 @@ sig | VernacLocate of locatable | VernacRegister of lident * register_kind | VernacComments of comment list - | VernacStm of stm_vernac | VernacGoal of Constrexpr.constr_expr | VernacAbort of lident option | VernacAbortAll diff --git a/dev/doc/xml-protocol.md b/dev/doc/xml-protocol.md index 127b4a6d2..cf7d205d8 100644 --- a/dev/doc/xml-protocol.md +++ b/dev/doc/xml-protocol.md @@ -291,7 +291,10 @@ Pseudocode for listing all of the goals in order: `rev (flat_map fst background) ### <a name="command-status">**Status(force: bool)**</a> -CoqIDE typically sets `force` to `false`. +Returns information about the current proofs. CoqIDE typically sends this +message with `force = false` after each sentence, and with `force = true` if +the user wants to force the checking of all proofs (wheels button). In terms of +the STM API, `force` triggers a `Join`. ```html <call val="Status"><bool val="${force}"/></call> ``` diff --git a/ide/ide_slave.ml b/ide/ide_slave.ml index b11a11606..f00b1e142 100644 --- a/ide/ide_slave.ml +++ b/ide/ide_slave.ml @@ -413,6 +413,7 @@ let eval_call c = Interface.quit = (fun () -> quit := true); Interface.init = interruptible init; Interface.about = interruptible about; + Interface.wait = interruptible Stm.wait; Interface.interp = interruptible interp; Interface.handle_exn = handle_exn; Interface.stop_worker = Stm.stop_worker; diff --git a/ide/interface.mli b/ide/interface.mli index 1939a8427..a5d98946f 100644 --- a/ide/interface.mli +++ b/ide/interface.mli @@ -229,6 +229,9 @@ type print_ast_rty = Xml_datatype.xml type annotate_sty = string type annotate_rty = Xml_datatype.xml +type wait_sty = unit +type wait_rty = unit + type handler = { add : add_sty -> add_rty; edit_at : edit_at_sty -> edit_at_rty; @@ -248,6 +251,8 @@ type handler = { handle_exn : handle_exn_sty -> handle_exn_rty; init : init_sty -> init_rty; quit : quit_sty -> quit_rty; + (* for internal use (fake_id) only, do not use *) + wait : wait_sty -> wait_rty; (* Retrocompatibility stuff *) interp : interp_sty -> interp_rty; } diff --git a/ide/xmlprotocol.ml b/ide/xmlprotocol.ml index 4b521a968..b452b0a13 100644 --- a/ide/xmlprotocol.ml +++ b/ide/xmlprotocol.ml @@ -531,6 +531,7 @@ let set_options_sty_t : set_options_sty val_t = list_t (pair_t (list_t string_t) option_value_t) let mkcases_sty_t : mkcases_sty val_t = string_t let quit_sty_t : quit_sty val_t = unit_t +let wait_sty_t : wait_sty val_t = unit_t let about_sty_t : about_sty val_t = unit_t let init_sty_t : init_sty val_t = option_t string_t let interp_sty_t : interp_sty val_t = pair_t (pair_t bool_t bool_t) string_t @@ -555,6 +556,7 @@ let get_options_rty_t : get_options_rty val_t = let set_options_rty_t : set_options_rty val_t = unit_t let mkcases_rty_t : mkcases_rty val_t = list_t (list_t string_t) let quit_rty_t : quit_rty val_t = unit_t +let wait_rty_t : wait_rty val_t = unit_t let about_rty_t : about_rty val_t = coq_info_t let init_rty_t : init_rty val_t = state_id_t let interp_rty_t : interp_rty val_t = pair_t state_id_t (union_t string_t string_t) @@ -576,6 +578,7 @@ let calls = [| "SetOptions", ($)set_options_sty_t, ($)set_options_rty_t; "MkCases", ($)mkcases_sty_t, ($)mkcases_rty_t; "Quit", ($)quit_sty_t, ($)quit_rty_t; + "Wait", ($)wait_sty_t, ($)wait_rty_t; "About", ($)about_sty_t, ($)about_rty_t; "Init", ($)init_sty_t, ($)init_rty_t; "Interp", ($)interp_sty_t, ($)interp_rty_t; @@ -600,6 +603,8 @@ type 'a call = | About : about_sty -> about_rty call | Init : init_sty -> init_rty call | StopWorker : stop_worker_sty -> stop_worker_rty call + (* internal use (fake_ide) only, do not use *) + | Wait : wait_sty -> wait_rty call (* retrocompatibility *) | Interp : interp_sty -> interp_rty call | PrintAst : print_ast_sty -> print_ast_rty call @@ -618,12 +623,13 @@ let id_of_call : type a. a call -> int = function | SetOptions _ -> 9 | MkCases _ -> 10 | Quit _ -> 11 - | About _ -> 12 - | Init _ -> 13 - | Interp _ -> 14 - | StopWorker _ -> 15 - | PrintAst _ -> 16 - | Annotate _ -> 17 + | Wait _ -> 12 + | About _ -> 13 + | Init _ -> 14 + | Interp _ -> 15 + | StopWorker _ -> 16 + | PrintAst _ -> 17 + | Annotate _ -> 18 let str_of_call c = pi1 calls.(id_of_call c) @@ -643,6 +649,7 @@ let mkcases x : mkcases_rty call = MkCases x let search x : search_rty call = Search x let quit x : quit_rty call = Quit x let init x : init_rty call = Init x +let wait x : wait_rty call = Wait x let interp x : interp_rty call = Interp x let stop_worker x : stop_worker_rty call = StopWorker x let print_ast x : print_ast_rty call = PrintAst x @@ -664,6 +671,7 @@ let abstract_eval_call : type a. _ -> a call -> a value = fun handler c -> | SetOptions x -> mkGood (handler.set_options x) | MkCases x -> mkGood (handler.mkcases x) | Quit x -> mkGood (handler.quit x) + | Wait x -> mkGood (handler.wait x) | About x -> mkGood (handler.about x) | Init x -> mkGood (handler.init x) | Interp x -> mkGood (handler.interp x) @@ -688,6 +696,7 @@ let of_answer : type a. a call -> a value -> xml = function | SetOptions _ -> of_value (of_value_type set_options_rty_t) | MkCases _ -> of_value (of_value_type mkcases_rty_t ) | Quit _ -> of_value (of_value_type quit_rty_t ) + | Wait _ -> of_value (of_value_type wait_rty_t ) | About _ -> of_value (of_value_type about_rty_t ) | Init _ -> of_value (of_value_type init_rty_t ) | Interp _ -> of_value (of_value_type interp_rty_t ) @@ -711,6 +720,7 @@ let to_answer : type a. a call -> xml -> a value = function | SetOptions _ -> to_value (to_value_type set_options_rty_t) | MkCases _ -> to_value (to_value_type mkcases_rty_t ) | Quit _ -> to_value (to_value_type quit_rty_t ) + | Wait _ -> to_value (to_value_type wait_rty_t ) | About _ -> to_value (to_value_type about_rty_t ) | Init _ -> to_value (to_value_type init_rty_t ) | Interp _ -> to_value (to_value_type interp_rty_t ) @@ -733,6 +743,7 @@ let of_call : type a. a call -> xml = fun q -> | SetOptions x -> mkCall (of_value_type set_options_sty_t x) | MkCases x -> mkCall (of_value_type mkcases_sty_t x) | Quit x -> mkCall (of_value_type quit_sty_t x) + | Wait x -> mkCall (of_value_type wait_sty_t x) | About x -> mkCall (of_value_type about_sty_t x) | Init x -> mkCall (of_value_type init_sty_t x) | Interp x -> mkCall (of_value_type interp_sty_t x) @@ -756,6 +767,7 @@ let to_call : xml -> unknown_call = | "SetOptions" -> Unknown (SetOptions (mkCallArg set_options_sty_t a)) | "MkCases" -> Unknown (MkCases (mkCallArg mkcases_sty_t a)) | "Quit" -> Unknown (Quit (mkCallArg quit_sty_t a)) + | "Wait" -> Unknown (Wait (mkCallArg wait_sty_t a)) | "About" -> Unknown (About (mkCallArg about_sty_t a)) | "Init" -> Unknown (Init (mkCallArg init_sty_t a)) | "Interp" -> Unknown (Interp (mkCallArg interp_sty_t a)) @@ -786,6 +798,7 @@ let pr_full_value : type a. a call -> a value -> string = fun call value -> matc | SetOptions _ -> pr_value_gen (print set_options_rty_t) value | MkCases _ -> pr_value_gen (print mkcases_rty_t ) value | Quit _ -> pr_value_gen (print quit_rty_t ) value + | Wait _ -> pr_value_gen (print wait_rty_t ) value | About _ -> pr_value_gen (print about_rty_t ) value | Init _ -> pr_value_gen (print init_rty_t ) value | Interp _ -> pr_value_gen (print interp_rty_t ) value @@ -807,6 +820,7 @@ let pr_call : type a. a call -> string = fun call -> | SetOptions x -> return set_options_sty_t x | MkCases x -> return mkcases_sty_t x | Quit x -> return quit_sty_t x + | Wait x -> return wait_sty_t x | About x -> return about_sty_t x | Init x -> return init_sty_t x | Interp x -> return interp_sty_t x diff --git a/ide/xmlprotocol.mli b/ide/xmlprotocol.mli index d1c678b90..22117e35c 100644 --- a/ide/xmlprotocol.mli +++ b/ide/xmlprotocol.mli @@ -29,6 +29,8 @@ val set_options : set_options_sty -> set_options_rty call val quit : quit_sty -> quit_rty call val init : init_sty -> init_rty call val stop_worker : stop_worker_sty -> stop_worker_rty call +(* internal use (fake_ide) only, do not use *) +val wait : wait_sty -> wait_rty call (* retrocompatibility *) val interp : interp_sty -> interp_rty call val print_ast : print_ast_sty -> print_ast_rty call diff --git a/intf/vernacexpr.ml b/intf/vernacexpr.ml index 94e166f92..fb713d352 100644 --- a/intf/vernacexpr.ml +++ b/intf/vernacexpr.ml @@ -280,11 +280,6 @@ type bullet = | Star of int | Plus of int -(** {6 Types concerning Stm} *) -type stm_vernac = - | JoinDocument - | Wait - (** {6 Types concerning the module layer} *) (** Rigid / flexible module signature *) @@ -450,10 +445,6 @@ type vernac_expr = | VernacRegister of lident * register_kind | VernacComments of comment list - (* Stm backdoor: used in fake_id, will be removed when fake_ide - becomes aware of feedback about completed jobs. *) - | VernacStm of stm_vernac - (* Proof management *) | VernacGoal of constr_expr | VernacAbort of lident option @@ -504,16 +495,12 @@ type vernac_type = | VtProofStep of proof_step | VtProofMode of string | VtQuery of vernac_part_of_script * Feedback.route_id - | VtStm of vernac_control * vernac_part_of_script + | VtBack of vernac_part_of_script * Stateid.t | VtUnknown and vernac_qed_type = VtKeep | VtKeepAsAxiom | VtDrop (* Qed/Admitted, Abort *) and vernac_start = string * opacity_guarantee * Id.t list and vernac_sideff_type = Id.t list and vernac_part_of_script = bool -and vernac_control = - | VtWait - | VtJoinDocument - | VtBack of Stateid.t and opacity_guarantee = | GuaranteesOpacity (** Only generates opaque terms at [Qed] *) | Doesn'tGuaranteeOpacity (** May generate transparent terms even with [Qed].*) diff --git a/parsing/g_vernac.ml4 b/parsing/g_vernac.ml4 index de98415cd..0da22fd71 100644 --- a/parsing/g_vernac.ml4 +++ b/parsing/g_vernac.ml4 @@ -75,10 +75,6 @@ GEXTEND Gram | IDENT "Local"; v = vernac_poly -> VernacLocal (true, v) | IDENT "Global"; v = vernac_poly -> VernacLocal (false, v) - (* Stm backdoor *) - | IDENT "Stm"; IDENT "JoinDocument"; "." -> VernacStm JoinDocument - | IDENT "Stm"; IDENT "Wait"; "." -> VernacStm Wait - | v = vernac_poly -> v ] ] ; diff --git a/printing/ppvernac.ml b/printing/ppvernac.ml index 5e57f5de0..51c3eb212 100644 --- a/printing/ppvernac.ml +++ b/printing/ppvernac.ml @@ -524,12 +524,6 @@ open Decl_kinds | VernacLocal (local, v) -> return (pr_locality local ++ spc() ++ pr_vernac_body v) - (* Stm *) - | VernacStm JoinDocument -> - return (keyword "Stm JoinDocument") - | VernacStm Wait -> - return (keyword "Stm Wait") - (* Proof management *) | VernacAbortAll -> return (keyword "Abort All") diff --git a/stm/stm.ml b/stm/stm.ml index 3386044f2..e0064df9b 100644 --- a/stm/stm.ml +++ b/stm/stm.ml @@ -1049,7 +1049,7 @@ end = struct (* {{{ *) try match v with | VernacResetInitial -> - VtStm (VtBack Stateid.initial, true), VtNow + VtBack (true, Stateid.initial), VtNow | VernacResetName (_,name) -> let id = VCS.get_branch_pos (VCS.current_branch ()) in (try @@ -1057,20 +1057,20 @@ end = struct (* {{{ *) fold_until (fun b (id,_,label,_,_) -> if b then `Stop id else `Cont (List.mem name label)) false id in - VtStm (VtBack oid, true), VtNow + VtBack (true, oid), VtNow with Not_found -> - VtStm (VtBack id, true), VtNow) + VtBack (true, id), VtNow) | VernacBack n -> let id = VCS.get_branch_pos (VCS.current_branch ()) in let oid = fold_until (fun n (id,_,_,_,_) -> if Int.equal n 0 then `Stop id else `Cont (n-1)) n id in - VtStm (VtBack oid, true), VtNow + VtBack (true, oid), VtNow | VernacUndo n -> let id = VCS.get_branch_pos (VCS.current_branch ()) in let oid = fold_until (fun n (id,_,_,tactic,undo) -> let value = (if tactic then 1 else 0) - undo in if Int.equal n 0 then `Stop id else `Cont (n-value)) n id in - VtStm (VtBack oid, true), VtLater + VtBack (true, oid), VtLater | VernacUndoTo _ | VernacRestart as e -> let m = match e with VernacUndoTo m -> m | _ -> 0 in @@ -1087,16 +1087,16 @@ end = struct (* {{{ *) 0 id in let oid = fold_until (fun n (id,_,_,_,_) -> if Int.equal n 0 then `Stop id else `Cont (n-1)) (n-m-1) id in - VtStm (VtBack oid, true), VtLater + VtBack (true, oid), VtLater | VernacAbortAll -> let id = VCS.get_branch_pos (VCS.current_branch ()) in let oid = fold_until (fun () (id,vcs,_,_,_) -> match Vcs_.branches vcs with [_] -> `Stop id | _ -> `Cont ()) () id in - VtStm (VtBack oid, true), VtLater + VtBack (true, oid), VtLater | VernacBacktrack (id,_,_) | VernacBackTo id -> - VtStm (VtBack (Stateid.of_int id), not !Flags.batch_mode), VtNow + VtBack (not !Flags.batch_mode, Stateid.of_int id), VtNow | _ -> VtUnknown, VtNow with | Not_found -> @@ -2373,6 +2373,7 @@ let finish () = | _ -> () let wait () = + finish (); Slaves.wait_all_done (); VCS.print () @@ -2386,7 +2387,6 @@ let rec join_admitted_proofs id = | _ -> join_admitted_proofs view.next let join () = - finish (); wait (); stm_prerr_endline (fun () -> "Joining the environment"); Global.join_safe_environment (); @@ -2484,14 +2484,8 @@ let process_transaction ?(newtip=Stateid.fresh ()) stm_prerr_endline (fun () -> " classified as: " ^ string_of_vernac_classification c); match c with - (* Joining various parts of the document *) - | VtStm (VtJoinDocument, b), VtNow -> join (); `Ok - | VtStm (VtWait, b), VtNow -> finish (); wait (); `Ok - | VtStm ((VtJoinDocument|VtWait),_), VtLater -> - anomaly(str"classifier: join actions cannot be classified as VtLater.") - (* Back *) - | VtStm (VtBack oid, true), w -> + | VtBack (true, oid), w -> let id = VCS.new_node ~id:newtip () in let { mine; others } = Backtrack.branches_of oid in let valid = VCS.get_branch_pos head in @@ -2510,12 +2504,12 @@ let process_transaction ?(newtip=Stateid.fresh ()) VCS.checkout_shallowest_proof_branch (); VCS.commit id (Alias (oid,x)); Backtrack.record (); if w == VtNow then finish (); `Ok - | VtStm (VtBack id, false), VtNow -> + | VtBack (false, id), VtNow -> stm_prerr_endline (fun () -> "undo to state " ^ Stateid.to_string id); Backtrack.backto id; VCS.checkout_shallowest_proof_branch (); Reach.known_state ~cache:(interactive ()) id; `Ok - | VtStm (VtBack id, false), VtLater -> + | VtBack (false, id), VtLater -> anomaly(str"classifier: VtBack + VtLater must imply part_of_script.") (* Query *) @@ -2779,8 +2773,8 @@ let query ~at ~route s = let clas = classify_vernac ast in let aast = { verbose = true; indentation; strlen; loc; expr = ast } in match clas with - | VtStm (w,_), _ -> - ignore(process_transaction aast (VtStm (w,false), VtNow)) + | VtBack (_,id), _ -> (* TODO: can this still happen ? *) + ignore(process_transaction aast (VtBack (false,id), VtNow)) | _ -> ignore(process_transaction aast (VtQuery (false, route), VtNow))) s diff --git a/stm/stm.mli b/stm/stm.mli index 188b176ba..3f01fca01 100644 --- a/stm/stm.mli +++ b/stm/stm.mli @@ -51,6 +51,9 @@ val edit_at : Stateid.t -> [ `NewTip | `Focus of focus ] (* Evaluates the tip of the current branch *) val finish : unit -> unit +(* Internal use (fake_ide) only, do not use *) +val wait : unit -> unit + val observe : Stateid.t -> unit val stop_worker : string -> unit diff --git a/stm/vernac_classifier.ml b/stm/vernac_classifier.ml index c2ebea961..158ad9084 100644 --- a/stm/vernac_classifier.ml +++ b/stm/vernac_classifier.ml @@ -31,8 +31,7 @@ let string_of_vernac_type = function Option.default "" proof_block_detection | VtProofMode s -> "ProofMode " ^ s | VtQuery (b, route) -> "Query " ^ string_of_in_script b ^ " route " ^ string_of_int route - | VtStm ((VtJoinDocument|VtWait), b) -> "Stm " ^ string_of_in_script b - | VtStm (VtBack _, b) -> "Stm Back " ^ string_of_in_script b + | VtBack (b, _) -> "Stm Back " ^ string_of_in_script b let string_of_vernac_when = function | VtLater -> "Later" @@ -64,9 +63,6 @@ let rec classify_vernac e = * look at the entire dag to detect this option. *) | VernacSetOption (["Universe"; "Polymorphism"],_) | VernacUnsetOption (["Universe"; "Polymorphism"]) -> VtSideff [], VtNow - (* Stm *) - | VernacStm Wait -> VtStm (VtWait, true), VtNow - | VernacStm JoinDocument -> VtStm (VtJoinDocument, true), VtNow (* Nested vernac exprs *) | VernacProgram e -> classify_vernac e | VernacLocal (_,e) -> classify_vernac e @@ -79,7 +75,7 @@ let rec 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 _ | VtProofMode _ ), _ as x -> x + | VtBack _ | VtProofMode _ ), _ as x -> x | VtQed _, _ -> VtProofStep { parallel = `No; proof_block_detection = None }, VtNow diff --git a/tools/fake_ide.ml b/tools/fake_ide.ml index a9da27ba2..79723431c 100644 --- a/tools/fake_ide.ml +++ b/tools/fake_ide.ml @@ -252,11 +252,9 @@ let eval_print l coq = let to_id, _ = get_id id in eval_call (query (0,(phrase, to_id))) coq | [ Tok(_,"WAIT") ] -> - let phrase = "Stm Wait." in - eval_call (query (0,(phrase,tip_id()))) coq + eval_call (wait ()) coq | [ Tok(_,"JOIN") ] -> - let phrase = "Stm JoinDocument." in - eval_call (query (0,(phrase,tip_id()))) coq + eval_call (status true) coq | [ Tok(_,"ASSERT"); Tok(_,"TIP"); Tok(_,id) ] -> let to_id, _ = get_id id in if not(Stateid.equal (Document.tip doc) to_id) then error "Wrong tip" diff --git a/toplevel/vernac.ml b/toplevel/vernac.ml index c33f6b9b8..4b97ee0dd 100644 --- a/toplevel/vernac.ml +++ b/toplevel/vernac.ml @@ -133,7 +133,7 @@ let rec interp_vernac sid (loc,com) = document. Hopefully this is fixed when VtBack can be removed and Undo etc... are just interpreted regularly. *) let is_proof_step = match fst (Vernac_classifier.classify_vernac v) with - | VtProofStep _ | VtStm (VtBack _, _) | VtStartProof _ -> true + | VtProofStep _ | VtBack (_, _) | VtStartProof _ -> true | _ -> false in diff --git a/vernac/vernacentries.ml b/vernac/vernacentries.ml index b0e438a8e..c7b8def0e 100644 --- a/vernac/vernacentries.ml +++ b/vernac/vernacentries.ml @@ -1923,7 +1923,6 @@ let interp ?proof ?loc locality poly c = | VernacTime _ -> assert false | VernacRedirect _ -> assert false | VernacTimeout _ -> assert false - | VernacStm _ -> assert false (* The STM should handle that, but LOAD bypasses the STM... *) | VernacAbortAll -> CErrors.user_err (str "AbortAll cannot be used through the Load command") @@ -2184,10 +2183,6 @@ let interp ?(verbosely=true) ?proof (loc,c) = let orig_program_mode = Flags.is_program_mode () in let rec aux ?locality ?polymorphism isprogcmd = function - (* This assert case will be removed when fake_ide can understand - completion feedback *) - | VernacStm _ -> assert false (* Done by Stm *) - | VernacProgram c when not isprogcmd -> aux ?locality ?polymorphism true c | VernacProgram _ -> user_err Pp.(str "Program mode specified twice") | VernacLocal (b, c) when Option.is_empty locality -> diff --git a/vernac/vernacprop.ml b/vernac/vernacprop.ml index fc11bcf4a..3cff1f14c 100644 --- a/vernac/vernacprop.ml +++ b/vernac/vernacprop.ml @@ -17,8 +17,7 @@ let rec is_navigation_vernac = function | VernacResetName _ | VernacBacktrack _ | VernacBackTo _ - | VernacBack _ - | VernacStm _ -> true + | VernacBack _ -> true | VernacRedirect (_, (_,c)) | VernacTime (_,c) -> is_navigation_vernac c (* Time Back* is harmless *) |