Source file vernacinterp.ml
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
130
131
132
133
134
135
136
137
138
139
140
141
142
143
144
145
146
147
148
149
150
151
152
153
154
155
156
157
158
159
160
161
162
163
164
165
166
167
168
169
170
171
172
173
174
175
176
177
178
179
180
181
182
183
184
185
186
187
188
189
190
191
192
193
194
195
196
197
198
199
200
201
202
203
204
205
206
207
208
209
210
211
212
213
214
215
216
217
218
219
220
221
222
223
224
225
226
227
228
229
230
231
232
233
234
235
236
237
238
239
240
241
242
243
244
245
246
247
248
249
250
251
252
253
254
255
256
257
258
259
260
261
262
263
264
265
266
267
268
269
270
271
272
273
open Vernacexpr
let vernac_pperr_endline = CDebug.create ~name:"vernacinterp" ()
let interp_typed_vernac (Vernacextend.TypedVernac { inprog; outprog; inproof; outproof; run })
~pm ~stack =
let open Vernacextend in
let module LStack = Vernacstate.LemmaStack in
let proof = Option.map LStack.get_top stack in
let pm', proof' = run
~pm:(InProg.cast (NeList.head pm) inprog)
~proof:(InProof.cast proof inproof)
in
let pm = OutProg.cast pm' outprog pm in
let stack = let open OutProof in
match stack, outproof with
| stack, No -> stack
| None, Close -> assert false
| Some stack, Close -> snd (LStack.pop stack)
| None, Update -> assert false
| Some stack, Update -> Some (LStack.map_top ~f:(fun _ -> proof') stack)
| stack, New -> Some (LStack.push stack proof')
in
stack, pm
let proof_mode_opt_name = ["Default";"Proof";"Mode"]
let get_default_proof_mode =
Goptions.declare_interpreted_string_option_and_ref
~depr:false
~key:proof_mode_opt_name
~value:(Pvernac.register_proof_mode "Noedit" Pvernac.Vernac_.noedit_mode)
(fun name -> match Pvernac.lookup_proof_mode name with
| Some pm -> pm
| None -> CErrors.user_err Pp.(str (Format.sprintf "No proof mode named \"%s\"." name)))
Pvernac.proof_mode_to_string
(** A global default timeout, controlled by option "Set Default Timeout n".
Use "Unset Default Timeout" to deactivate it (or set it to 0). *)
let default_timeout = ref None
let vernac_timeout ?timeout (f : 'a -> 'b) (x : 'a) : 'b =
match !default_timeout, timeout with
| _, Some n
| Some n, None ->
(match Control.timeout (float_of_int n) f x with
| None -> Exninfo.iraise (Exninfo.capture CErrors.Timeout)
| Some x -> x)
| None, None ->
f x
let test_mode = ref false
let with_fail f : (Loc.t option * Pp.t, unit) result =
try
let _ = f () in
Error ()
with
| e ->
let _, info as exn = Exninfo.capture e in
if CErrors.is_anomaly e && e != CErrors.Timeout then Exninfo.iraise exn;
Ok (Loc.get_loc info, CErrors.iprint exn)
let real_error_loc ~cmdloc ~eloc =
if Loc.finer eloc cmdloc then eloc
else cmdloc
let with_fail ~loc ~st f =
let res = with_fail f in
Vernacstate.invalidate_cache ();
Vernacstate.unfreeze_interp_state st;
match res with
| Error () ->
CErrors.user_err (Pp.str "The command has not failed!")
| Ok (eloc, msg) ->
let loc = if !test_mode then real_error_loc ~cmdloc:loc ~eloc else None in
if not !Flags.quiet || !test_mode
then Feedback.msg_notice ?loc Pp.(str "The command has indeed failed with message:" ++ fnl () ++ msg)
let with_succeed ~st f =
let () = ignore (f ()) in
Vernacstate.invalidate_cache ();
Vernacstate.unfreeze_interp_state st;
if not !Flags.quiet
then Feedback.msg_notice Pp.(str "The command has succeeded and its effects have been reverted.")
let locate_if_not_already ?loc (e, info) =
(e, Option.cata (Loc.add_loc info) info (real_error_loc ~cmdloc:loc ~eloc:(Loc.get_loc info)))
let =
let vernac =
let vernac = match vernac with
| { CAst.v = { control = ControlTime _ :: control; attrs; expr }; loc } ->
CAst.make ?loc { control; attrs; expr }
| _ -> vernac
in
Topfmt.pr_cmd_header vernac
in
fun vernac -> Lazy.from_fun (fun () -> pr_time_header vernac)
let interp_control_flag ~loc ~ (f : control_flag) ~st
(fn : st:Vernacstate.t -> Vernacstate.LemmaStack.t option * Declare.OblState.t NeList.t) =
match f with
| ControlFail ->
with_fail ~loc ~st (fun () -> fn ~st);
st.Vernacstate.lemmas, st.Vernacstate.program
| ControlSucceed ->
with_succeed ~st (fun () -> fn ~st);
st.Vernacstate.lemmas, st.Vernacstate.program
| ControlTimeout timeout ->
vernac_timeout ~timeout (fun () -> fn ~st) ()
| ControlTime batch ->
let = if batch then Lazy.force time_header else Pp.mt () in
System.with_time ~batch ~header (fun () -> fn ~st) ()
| ControlRedirect s ->
Topfmt.with_output_to_file s (fun () -> fn ~st) ()
let rec interp_expr ?loc ~atts ~st c =
vernac_pperr_endline Pp.(fun () -> str "interpreting: " ++ Ppvernac.pr_vernac_expr c);
match c with
| VernacAbortAll -> CErrors.user_err (Pp.str "AbortAll cannot be used through the Load command")
| VernacRestart -> CErrors.user_err (Pp.str "Restart cannot be used through the Load command")
| VernacUndo _ -> CErrors.user_err (Pp.str "Undo cannot be used through the Load command")
| VernacUndoTo _ -> CErrors.user_err (Pp.str "UndoTo cannot be used through the Load command")
| VernacResetName _ -> CErrors.anomaly (Pp.str "VernacResetName not handled by Stm.")
| VernacResetInitial -> CErrors.anomaly (Pp.str "VernacResetInitial not handled by Stm.")
| VernacBack _ -> CErrors.anomaly (Pp.str "VernacBack not handled by Stm.")
| VernacLoad (verbosely, fname) ->
Attributes.unsupported_attributes atts;
vernac_load ~verbosely fname
| v ->
let fv = Vernacentries.translate_vernac ?loc ~atts v in
let stack = st.Vernacstate.lemmas in
let program = st.Vernacstate.program in
interp_typed_vernac ~pm:program ~stack fv
and vernac_load ~verbosely fname =
let st = Vernacstate.freeze_interp_state ~marshallable:false in
let fname =
Envars.expand_path_macros ~warn:(fun x -> Feedback.msg_warning (Pp.str x)) fname in
let fname = CUnix.make_suffix fname ".v" in
let input =
let longfname = Loadpath.locate_file fname in
let in_chan = Util.open_utf8_file_in longfname in
Pcoq.Parsable.make ~loc:Loc.(initial (InFile { dirpath=None; file=longfname}))
(Gramlib.Stream.of_channel in_chan) in
let v_mod = if verbosely then Flags.verbosely else Flags.silently in
let parse_sentence proof_mode = Flags.with_option Flags.we_are_parsing
(Pcoq.Entry.parse (Pvernac.main_entry proof_mode))
in
let rec load_loop ~pm ~stack =
let proof_mode = Option.map (fun _ -> get_default_proof_mode ()) stack in
match parse_sentence proof_mode input with
| None -> stack, pm
| Some stm ->
let stack, pm = v_mod (interp_control ~st:{ st with Vernacstate.lemmas = stack; program = pm }) stm in
(load_loop [@ocaml.tailcall]) ~stack ~pm
in
let stack, pm =
Dumpglob.with_glob_output Dumpglob.NoGlob
(fun () -> load_loop ~pm:st.Vernacstate.program ~stack:st.Vernacstate.lemmas) ()
in
if Option.has_some stack then
CErrors.user_err Pp.(str "Files processed by Load cannot leave open proofs.");
stack, pm
and interp_control ~st ({ CAst.v = cmd; loc } as vernac) =
let = mk_time_header vernac in
List.fold_right (fun flag fn -> interp_control_flag ~loc ~time_header flag fn)
cmd.control
(fun ~st ->
let before_univs = Global.universes () in
let pstack, pm = interp_expr ?loc ~atts:cmd.attrs ~st cmd.expr in
let after_univs = Global.universes () in
if before_univs == after_univs then pstack, pm
else
let f = Declare.Proof.update_sigma_univs after_univs in
Option.map (Vernacstate.LemmaStack.map ~f) pstack, pm)
~st
let interp_qed_delayed ~proof ~st pe =
let stack = st.Vernacstate.lemmas in
let pm = st.Vernacstate.program in
let stack = Option.cata (fun stack -> snd @@ Vernacstate.LemmaStack.pop stack) None stack in
let pm = NeList.map_head (fun pm -> match pe with
| Admitted ->
Declare.Proof.save_lemma_admitted_delayed ~pm ~proof
| Proved (_,idopt) ->
let pm, _ = Declare.Proof.save_lemma_proved_delayed ~pm ~proof ~idopt in
pm)
pm
in
stack, pm
let interp_qed_delayed_control ~proof ~st ~control { CAst.loc; v=pe } =
let = mk_time_header (CAst.make ?loc { control; attrs = []; expr = VernacEndProof pe }) in
List.fold_right (fun flag fn -> interp_control_flag ~loc ~time_header flag fn)
control
(fun ~st -> interp_qed_delayed ~proof ~st pe)
~st
let () = let open Goptions in
declare_int_option
{ optdepr = false;
optkey = ["Default";"Timeout"];
optread = (fun () -> !default_timeout);
optwrite = ((:=) default_timeout) }
let interp_gen ~verbosely ~st ~interp_fn cmd =
Vernacstate.unfreeze_interp_state st;
try vernac_timeout (fun st ->
let v_mod = if verbosely then Flags.verbosely else Flags.silently in
let ontop = v_mod (interp_fn ~st) cmd in
Vernacstate.Declare.set ontop [@ocaml.warning "-3"];
Vernacstate.freeze_interp_state ~marshallable:false
) st
with exn ->
let exn = Exninfo.capture exn in
let exn = locate_if_not_already ?loc:cmd.CAst.loc exn in
Vernacstate.invalidate_cache ();
Exninfo.iraise exn
let interp ?(verbosely=true) ~st cmd =
interp_gen ~verbosely ~st ~interp_fn:interp_control cmd
let interp_qed_delayed_proof ~proof ~st ~control pe : Vernacstate.t =
interp_gen ~verbosely:false ~st
~interp_fn:(interp_qed_delayed_control ~proof ~control) pe