Skip to content

Commit

Permalink
Add an option not to register Ltac backtraces.
Browse files Browse the repository at this point in the history
Fixes coq#7769: Better control over ltac backtraces.
Fixes coq#7385: Tactic allocates gigabytes of RAM.

For the lulz I disabled the option by default, let's see if that makes
any performance difference.
  • Loading branch information
ppedrot committed Dec 5, 2018
1 parent 23f2222 commit 1f8cbc1
Show file tree
Hide file tree
Showing 2 changed files with 31 additions and 11 deletions.
2 changes: 2 additions & 0 deletions CHANGES.md
Expand Up @@ -72,6 +72,8 @@ Tactics
foo : database`). When the database name is omitted, the hint is added to the
core database (as previously), but a deprecation warning is emitted.

- Ltac backtraces can be turned off using the "Ltac Backtrace" option.

Vernacular commands

- `Combined Scheme` can now work when inductive schemes are generated in sort
Expand Down
40 changes: 29 additions & 11 deletions plugins/ltac/tacinterp.ml
Expand Up @@ -117,9 +117,11 @@ let combine_appl appl1 appl2 =
let of_tacvalue v = in_gen (topwit wit_tacvalue) v
let to_tacvalue v = out_gen (topwit wit_tacvalue) v

let log_trace = ref false

(** More naming applications *)
let name_vfun appl vle =
if has_type vle (topwit wit_tacvalue) then
if !log_trace && has_type vle (topwit wit_tacvalue) then
match to_tacvalue vle with
| VFun (appl0,trace,lfun,vars,t) -> of_tacvalue (VFun (combine_appl appl0 appl,trace,lfun,vars,t))
| _ -> vle
Expand All @@ -137,9 +139,11 @@ type interp_sign = Geninterp.interp_sign = {
lfun : value Id.Map.t;
extra : TacStore.t }

let extract_trace ist = match TacStore.get ist.extra f_trace with
| None -> []
| Some l -> l
let extract_trace ist =
if !log_trace then match TacStore.get ist.extra f_trace with
| None -> []
| Some l -> l
else []

let print_top_val env v = Pptactic.pr_value Pptactic.ltop v

Expand All @@ -161,8 +165,11 @@ let catch_error call_trace f x =
let e = CErrors.push e in
catching_error call_trace iraise e

let wrap_error tac k =
if !log_trace then Proofview.tclORELSE tac k else tac

let catch_error_tac call_trace tac =
Proofview.tclORELSE
wrap_error
tac
(catching_error call_trace (fun (e, info) -> Proofview.tclZERO ~info e))

Expand Down Expand Up @@ -203,9 +210,11 @@ let constr_of_id env id =
(** Generic arguments : table of interpretation functions *)

(* Some of the code further down depends on the fact that push_trace does not modify sigma (the evar map) *)
let push_trace call ist = match TacStore.get ist.extra f_trace with
| None -> Proofview.tclUNIT [call]
| Some trace -> Proofview.tclUNIT (call :: trace)
let push_trace call ist =
if !log_trace then match TacStore.get ist.extra f_trace with
| None -> Proofview.tclUNIT [call]
| Some trace -> Proofview.tclUNIT (call :: trace)
else Proofview.tclUNIT []

let propagate_trace ist loc id v =
if has_type v (topwit wit_tacvalue) then
Expand Down Expand Up @@ -1247,7 +1256,7 @@ and interp_app loc ist fv largs : Val.t Ftactic.t =
let fold accu (id, v) = Id.Map.add id v accu in
let newlfun = List.fold_left fold olfun extfun in
if List.is_empty lvar then
begin Proofview.tclORELSE
begin wrap_error
begin
let ist = {
lfun = newlfun;
Expand Down Expand Up @@ -1407,7 +1416,7 @@ and interp_match_successes lz ist s =
(* Interprets the Match expressions *)
and interp_match ist lz constr lmr =
let (>>=) = Ftactic.bind in
begin Proofview.tclORELSE
begin wrap_error
(interp_ltac_constr ist constr)
begin function
| (e, info) ->
Expand Down Expand Up @@ -1493,7 +1502,7 @@ and interp_genarg_var_list ist x =
(* Interprets tactic expressions : returns a "constr" *)
and interp_ltac_constr ist e : EConstr.t Ftactic.t =
let (>>=) = Ftactic.bind in
begin Proofview.tclORELSE
begin wrap_error
(val_interp ist e)
begin function (err, info) -> match err with
| Not_found ->
Expand Down Expand Up @@ -2048,4 +2057,13 @@ let () =
optread = (fun () -> get_debug () != Tactic_debug.DebugOff);
optwrite = vernac_debug }

let () =
let open Goptions in
declare_bool_option
{ optdepr = false;
optname = "Ltac Backtrace";
optkey = ["Ltac"; "Backtrace"];
optread = (fun () -> !log_trace);
optwrite = (fun b -> log_trace := b) }

let () = Hook.set Vernacentries.interp_redexp_hook interp_redexp

0 comments on commit 1f8cbc1

Please sign in to comment.