| author | wenzelm | 
| Sat, 13 Aug 2022 14:45:36 +0200 | |
| changeset 75841 | 7c00d5266bf8 | 
| parent 75080 | 1dae5cbcd358 | 
| child 76183 | 8089593a364a | 
| permissions | -rw-r--r-- | 
| 47847 | 1 | (* Title: HOL/Mirabelle/Tools/mirabelle.ML | 
| 73847 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 2 | Author: Jasmin Blanchette, TU Munich | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 3 | Author: Sascha Boehme, TU Munich | 
| 73691 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 4 | Author: Makarius | 
| 73847 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 5 | Author: Martin Desharnais, UniBw Munich | 
| 32381 | 6 | *) | 
| 7 | ||
| 8 | signature MIRABELLE = | |
| 9 | sig | |
| 32564 | 10 | (*core*) | 
| 74078 
a2cbe81e1e32
changed Mirabelle_Sledgehammer keep option from path to boolean
 desharna parents: 
74077diff
changeset | 11 | type action_context = | 
| 74125 | 12 |     {index: int, label: string, name: string, arguments: Properties.T, timeout: Time.time,
 | 
| 13 | output_dir: Path.T} | |
| 73847 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 14 | type command = | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 15 |     {theory_index: int, name: string, pos: Position.T, pre: Proof.state, post: Toplevel.state}
 | 
| 75003 
f21e7e6172a0
renamed run_action to run in Mirabelle.action record
 desharna parents: 
74991diff
changeset | 16 |   type action = {run: command -> string, finalize: unit -> string}
 | 
| 74948 
15ce207f69c8
added support for initialization messages to Mirabelle
 desharna parents: 
74745diff
changeset | 17 | val register_action: string -> (action_context -> string * action) -> unit | 
| 32495 | 18 | |
| 32564 | 19 | (*utility functions*) | 
| 73847 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 20 | val print_exn: exn -> string | 
| 32469 | 21 | val can_apply : Time.time -> (Proof.context -> int -> tactic) -> | 
| 22 | Proof.state -> bool | |
| 60361 | 23 | val theorems_in_proof_term : theory -> thm -> thm list | 
| 73691 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 24 | val theorems_of_sucessful_proof: Toplevel.state -> thm list | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 25 | val get_argument : (string * string) list -> string * string -> string | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 26 | val get_int_argument : (string * string) list -> string * int -> int | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 27 | val get_bool_argument : (string * string) list -> string * bool -> bool | 
| 32498 
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
 boehmes parents: 
32497diff
changeset | 28 |   val cpu_time : ('a -> 'b) -> 'a -> 'b * int
 | 
| 32381 | 29 | end | 
| 30 | ||
| 32497 | 31 | structure Mirabelle : MIRABELLE = | 
| 32495 | 32 | struct | 
| 33 | ||
| 73691 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 34 | (** Mirabelle core **) | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 35 | |
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 36 | (* concrete syntax *) | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 37 | |
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 38 | fun read_actions str = | 
| 74125 | 39 | let | 
| 74558 | 40 | val thy = \<^theory>; | 
| 41 | val ctxt = Proof_Context.init_global thy | |
| 74567 | 42 | val keywords = Keyword.no_major_keywords (Thy_Header.get_keywords thy) | 
| 74558 | 43 | |
| 74564 | 44 | fun read_actions () = Parse.read_embedded ctxt keywords | 
| 74125 | 45 | (Parse.enum ";" (Parse.name -- Sledgehammer_Commands.parse_params)) | 
| 74558 | 46 | (Input.string str) | 
| 74125 | 47 | fun split_name_label (name, args) labels = | 
| 48 | (case String.tokens (fn c => c = #".") name of | |
| 49 |           [name] => (("", name, args), labels)
 | |
| 50 | | [label, name] => | |
| 51 | (if Symtab.defined labels label then | |
| 52 |              error ("Action label '" ^ label ^ "' already defined.")
 | |
| 53 | else | |
| 54 | (); | |
| 55 | ((label, name, args), Symtab.insert_set label labels)) | |
| 56 | | _ => error "Cannot parse action") | |
| 57 | in | |
| 74558 | 58 | try read_actions () | 
| 59 | |> Option.map (fn xs => fst (fold_map split_name_label xs Symtab.empty)) | |
| 74125 | 60 | end | 
| 32385 
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
 boehmes parents: 
32383diff
changeset | 61 | |
| 
594890623c46
split actions from Mirabelle core (Mirabelle may thus be applied to basic theories in HOL)
 boehmes parents: 
32383diff
changeset | 62 | |
| 73691 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 63 | (* actions *) | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 64 | |
| 73847 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 65 | type command = | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 66 |   {theory_index: int, name: string, pos: Position.T, pre: Proof.state, post: Toplevel.state};
 | 
| 74078 
a2cbe81e1e32
changed Mirabelle_Sledgehammer keep option from path to boolean
 desharna parents: 
74077diff
changeset | 67 | type action_context = | 
| 74125 | 68 |   {index: int, label: string, name: string, arguments: Properties.T, timeout: Time.time,
 | 
| 69 | output_dir: Path.T}; | |
| 75003 
f21e7e6172a0
renamed run_action to run in Mirabelle.action record
 desharna parents: 
74991diff
changeset | 70 | type action = {run: command -> string, finalize: unit -> string};
 | 
| 32381 | 71 | |
| 75003 
f21e7e6172a0
renamed run_action to run in Mirabelle.action record
 desharna parents: 
74991diff
changeset | 72 | val dry_run_action : action = {run = K "", finalize = K ""}
 | 
| 74989 | 73 | |
| 73847 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 74 | local | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 75 | val actions = Synchronized.var "Mirabelle.actions" | 
| 74948 
15ce207f69c8
added support for initialization messages to Mirabelle
 desharna parents: 
74745diff
changeset | 76 | (Symtab.empty : (action_context -> string * action) Symtab.table); | 
| 73847 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 77 | in | 
| 73694 | 78 | |
| 73851 
bb277f37c34a
added warnings when defining unamed or redefining Mirabelle action
 desharna parents: 
73850diff
changeset | 79 | fun register_action name make_action = | 
| 
bb277f37c34a
added warnings when defining unamed or redefining Mirabelle action
 desharna parents: 
73850diff
changeset | 80 | (if name = "" then error "Registering unnamed Mirabelle action" else (); | 
| 
bb277f37c34a
added warnings when defining unamed or redefining Mirabelle action
 desharna parents: 
73850diff
changeset | 81 | Synchronized.change actions (Symtab.map_default (name, make_action) | 
| 
bb277f37c34a
added warnings when defining unamed or redefining Mirabelle action
 desharna parents: 
73850diff
changeset | 82 |      (fn f => (warning ("Redefining Mirabelle action: " ^ quote name); f))));
 | 
| 73694 | 83 | |
| 73847 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 84 | fun get_action name = Symtab.lookup (Synchronized.value actions) name; | 
| 73694 | 85 | |
| 73847 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 86 | end | 
| 73694 | 87 | |
| 88 | ||
| 89 | (* apply actions *) | |
| 90 | ||
| 73691 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 91 | fun print_exn exn = | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 92 | (case exn of | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 93 | Timeout.TIMEOUT _ => "timeout" | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 94 | | ERROR msg => "error: " ^ msg | 
| 73847 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 95 | | exn => "exception: " ^ General.exnMessage exn); | 
| 34035 
08d34921b7dd
also consider the fully-typed version of metis for Mirabelle measurements
 boehmes parents: 
33522diff
changeset | 96 | |
| 73847 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 97 | fun run_action_function f = | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 98 | f () handle exn => | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 99 | if Exn.is_interrupt exn then Exn.reraise exn | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 100 | else print_exn exn; | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 101 | |
| 74745 | 102 | fun make_action_path ({index, label, name, ...} : action_context) =
 | 
| 103 | Path.basic (if label = "" then string_of_int index ^ "." ^ name else label); | |
| 73847 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 104 | |
| 74948 
15ce207f69c8
added support for initialization messages to Mirabelle
 desharna parents: 
74745diff
changeset | 105 | fun initialize_action (make_action : action_context -> string * action) context = | 
| 
15ce207f69c8
added support for initialization messages to Mirabelle
 desharna parents: 
74745diff
changeset | 106 | let | 
| 
15ce207f69c8
added support for initialization messages to Mirabelle
 desharna parents: 
74745diff
changeset | 107 | val (s, action) = make_action context | 
| 
15ce207f69c8
added support for initialization messages to Mirabelle
 desharna parents: 
74745diff
changeset | 108 | val action_path = make_action_path context; | 
| 
15ce207f69c8
added support for initialization messages to Mirabelle
 desharna parents: 
74745diff
changeset | 109 | val export_name = | 
| 
15ce207f69c8
added support for initialization messages to Mirabelle
 desharna parents: 
74745diff
changeset | 110 | Path.binding0 (Path.basic "mirabelle" + action_path + Path.basic "initialize"); | 
| 
15ce207f69c8
added support for initialization messages to Mirabelle
 desharna parents: 
74745diff
changeset | 111 | val () = | 
| 
15ce207f69c8
added support for initialization messages to Mirabelle
 desharna parents: 
74745diff
changeset | 112 | if s <> "" then | 
| 
15ce207f69c8
added support for initialization messages to Mirabelle
 desharna parents: 
74745diff
changeset | 113 | Export.export \<^theory> export_name [XML.Text s] | 
| 
15ce207f69c8
added support for initialization messages to Mirabelle
 desharna parents: 
74745diff
changeset | 114 | else | 
| 
15ce207f69c8
added support for initialization messages to Mirabelle
 desharna parents: 
74745diff
changeset | 115 | () | 
| 
15ce207f69c8
added support for initialization messages to Mirabelle
 desharna parents: 
74745diff
changeset | 116 | in | 
| 
15ce207f69c8
added support for initialization messages to Mirabelle
 desharna parents: 
74745diff
changeset | 117 | action | 
| 
15ce207f69c8
added support for initialization messages to Mirabelle
 desharna parents: 
74745diff
changeset | 118 | end | 
| 
15ce207f69c8
added support for initialization messages to Mirabelle
 desharna parents: 
74745diff
changeset | 119 | |
| 73847 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 120 | fun finalize_action ({finalize, ...} : action) context =
 | 
| 73691 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 121 | let | 
| 73847 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 122 | val s = run_action_function finalize; | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 123 | val action_path = make_action_path context; | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 124 | val export_name = | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 125 | Path.binding0 (Path.basic "mirabelle" + action_path + Path.basic "finalize"); | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 126 | in | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 127 | if s <> "" then | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 128 | Export.export \<^theory> export_name [XML.Text s] | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 129 | else | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 130 | () | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 131 | end | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 132 | |
| 75003 
f21e7e6172a0
renamed run_action to run in Mirabelle.action record
 desharna parents: 
74991diff
changeset | 133 | fun apply_action ({run, ...} : action) context (command as {pos, pre, ...} : command) =
 | 
| 73847 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 134 | let | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 135 | val thy = Proof.theory_of pre; | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 136 | val action_path = make_action_path context; | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 137 | val goal_name_path = Path.basic (#name command) | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 138 | val line_path = Path.basic (string_of_int (the (Position.line_of pos))); | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 139 | val offset_path = Path.basic (string_of_int (the (Position.offset_of pos))); | 
| 75003 
f21e7e6172a0
renamed run_action to run in Mirabelle.action record
 desharna parents: 
74991diff
changeset | 140 |     val ({elapsed, ...}, s) = Timing.timing run_action_function (fn () => run command);
 | 
| 73847 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 141 | val export_name = | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 142 | Path.binding0 (Path.basic "mirabelle" + action_path + Path.basic "goal" + goal_name_path + | 
| 74991 
d699eb2d26ad
used elapsed time instead of cpu time in Mirabelle because the latter contain cpu time of all threads
 desharna parents: 
74989diff
changeset | 143 | line_path + offset_path + Path.basic (string_of_int (Time.toMilliseconds elapsed))); | 
| 73847 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 144 | in | 
| 74989 | 145 | Export.export thy export_name [XML.Text s] | 
| 73847 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 146 | end; | 
| 32515 
e7c0d3c0494a
Mirabelle: actions are responsible for catching exceptions and producing suitable log messages (makes log message uniform),
 boehmes parents: 
32504diff
changeset | 147 | |
| 73691 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 148 | |
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 149 | (* theory line range *) | 
| 32381 | 150 | |
| 151 | local | |
| 152 | ||
| 73691 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 153 | val theory_name = | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 154 | Scan.many1 (Symbol_Pos.symbol #> (fn s => Symbol.not_eof s andalso s <> "[")) | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 155 | >> Symbol_Pos.content; | 
| 32498 
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
 boehmes parents: 
32497diff
changeset | 156 | |
| 73691 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 157 | val line = Symbol_Pos.scan_nat >> (Symbol_Pos.content #> Value.parse_nat); | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 158 | val end_line = Symbol_Pos.$$ ":" |-- line; | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 159 | val range = Symbol_Pos.$$ "[" |-- line -- Scan.option end_line --| Symbol_Pos.$$ "]"; | 
| 32381 | 160 | |
| 161 | in | |
| 162 | ||
| 73691 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 163 | fun read_theory_range str = | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 164 | (case Scan.read Symbol_Pos.stopper (theory_name -- Scan.option range) (Symbol_Pos.explode0 str) of | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 165 | SOME res => res | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 166 |   | NONE => error ("Malformed specification of theory line range: " ^ quote str));
 | 
| 32472 
7b92a8b8daaf
Mirabelle: actions are responsible for their log messages, output is better readable
 boehmes parents: 
32469diff
changeset | 167 | |
| 73691 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 168 | end; | 
| 32381 | 169 | |
| 73691 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 170 | fun check_theories strs = | 
| 32521 | 171 | let | 
| 74069 | 172 | fun theory_import_name s = | 
| 173 | #theory_name (Resources.import_name (Session.get_name ()) Path.current s); | |
| 174 | val theories = map read_theory_range strs | |
| 175 | |> map (apfst theory_import_name); | |
| 73691 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 176 | fun get_theory name = | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 177 | if null theories then SOME NONE | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 178 | else get_first (fn (a, b) => if a = name then SOME b else NONE) theories; | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 179 | fun check_line NONE _ = false | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 180 | | check_line _ NONE = true | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 181 | | check_line (SOME NONE) _ = true | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 182 | | check_line (SOME (SOME (line, NONE))) (SOME i) = line <= i | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 183 | | check_line (SOME (SOME (line, SOME end_line))) (SOME i) = line <= i andalso i <= end_line; | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 184 | fun check_pos range = check_line range o Position.line_of; | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 185 | in check_pos o get_theory end; | 
| 32521 | 186 | |
| 73691 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 187 | |
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 188 | (* presentation hook *) | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 189 | |
| 75080 
1dae5cbcd358
Mirabelle now considers goals preceding "unfolding" and "using" commands
 desharna parents: 
75003diff
changeset | 190 | val whitelist = ["apply", "by", "proof", "unfolding", "using"]; | 
| 32381 | 191 | |
| 73691 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 192 | val _ = | 
| 73822 
1192c68ebe1c
suppress theories from other sessions, unless explicitly specified via mirabelle_theories;
 wenzelm parents: 
73821diff
changeset | 193 | Build.add_hook (fn qualifier => fn loaded_theories => | 
| 73691 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 194 | let | 
| 73820 
745e2cd1f5f5
refer to theory "segments" only, according to global Build.build_theories and Thy_Info.use_theories;
 wenzelm parents: 
73808diff
changeset | 195 | val mirabelle_actions = Options.default_string \<^system_option>\<open>mirabelle_actions\<close>; | 
| 73691 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 196 | val actions = | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 197 | (case read_actions mirabelle_actions of | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 198 | SOME actions => actions | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 199 |         | NONE => error ("Failed to parse mirabelle_actions: " ^ quote mirabelle_actions));
 | 
| 73847 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 200 | in | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 201 | if null actions then | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 202 | () | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 203 | else | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 204 | let | 
| 74989 | 205 | val mirabelle_dry_run = Options.default_bool \<^system_option>\<open>mirabelle_dry_run\<close>; | 
| 73849 | 206 | val mirabelle_timeout = Options.default_seconds \<^system_option>\<open>mirabelle_timeout\<close>; | 
| 207 | val mirabelle_stride = Options.default_int \<^system_option>\<open>mirabelle_stride\<close>; | |
| 208 | val mirabelle_max_calls = Options.default_int \<^system_option>\<open>mirabelle_max_calls\<close>; | |
| 74986 
fc664e4fbf6d
added Mirabelle option -r to randomize the goals before selection
 desharna parents: 
74948diff
changeset | 209 | val mirabelle_randomize = Options.default_int \<^system_option>\<open>mirabelle_randomize\<close>; | 
| 73849 | 210 | val mirabelle_theories = Options.default_string \<^system_option>\<open>mirabelle_theories\<close>; | 
| 74745 | 211 | val mirabelle_output_dir = Options.default_string \<^system_option>\<open>mirabelle_output_dir\<close>; | 
| 73847 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 212 | val check_theory = check_theories (space_explode "," mirabelle_theories); | 
| 32381 | 213 | |
| 73847 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 214 | fun make_commands (thy_index, (thy, segments)) = | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 215 | let | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 216 | val thy_long_name = Context.theory_long_name thy; | 
| 74069 | 217 | val check_thy = check_theory thy_long_name; | 
| 73847 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 218 |               fun make_command {command = tr, prev_state = st, state = st', ...} =
 | 
| 73821 
9ead8d9be3ab
clarified hook for Mirabelle: provide all loaded theories at once (for each 'theories' section within the session ROOT);
 wenzelm parents: 
73820diff
changeset | 219 | let | 
| 
9ead8d9be3ab
clarified hook for Mirabelle: provide all loaded theories at once (for each 'theories' section within the session ROOT);
 wenzelm parents: 
73820diff
changeset | 220 | val name = Toplevel.name_of tr; | 
| 
9ead8d9be3ab
clarified hook for Mirabelle: provide all loaded theories at once (for each 'theories' section within the session ROOT);
 wenzelm parents: 
73820diff
changeset | 221 | val pos = Toplevel.pos_of tr; | 
| 
9ead8d9be3ab
clarified hook for Mirabelle: provide all loaded theories at once (for each 'theories' section within the session ROOT);
 wenzelm parents: 
73820diff
changeset | 222 | in | 
| 73824 | 223 | if Context.proper_subthy (\<^theory>, thy) andalso | 
| 224 | can (Proof.assert_backward o Toplevel.proof_of) st andalso | |
| 73847 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 225 | member (op =) whitelist name andalso check_thy pos | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 226 |                   then SOME {theory_index = thy_index, name = name, pos = pos,
 | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 227 | pre = Toplevel.proof_of st, post = st'} | 
| 73821 
9ead8d9be3ab
clarified hook for Mirabelle: provide all loaded theories at once (for each 'theories' section within the session ROOT);
 wenzelm parents: 
73820diff
changeset | 228 | else NONE | 
| 73847 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 229 | end; | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 230 | in | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 231 | if Resources.theory_qualifier thy_long_name = qualifier then | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 232 | map_filter make_command segments | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 233 | else | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 234 | [] | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 235 | end; | 
| 73821 
9ead8d9be3ab
clarified hook for Mirabelle: provide all loaded theories at once (for each 'theories' section within the session ROOT);
 wenzelm parents: 
73820diff
changeset | 236 | |
| 73847 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 237 | (* initialize actions *) | 
| 74125 | 238 | val contexts = actions |> map_index (fn (n, (label, name, args)) => | 
| 73847 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 239 | let | 
| 74514 | 240 | val make_action = | 
| 241 | (case get_action name of | |
| 242 | SOME f => f | |
| 243 | | NONE => error "Unknown action"); | |
| 74511 
6d111935299c
produced Mirabelle output directly in ML until Scala output gets fixed
 desharna parents: 
74125diff
changeset | 244 | val action_subdir = if label = "" then string_of_int n ^ "." ^ name else label; | 
| 74745 | 245 | val output_dir = | 
| 246 | Path.append (Path.explode mirabelle_output_dir) (Path.basic action_subdir); | |
| 74078 
a2cbe81e1e32
changed Mirabelle_Sledgehammer keep option from path to boolean
 desharna parents: 
74077diff
changeset | 247 | val context = | 
| 74125 | 248 |                 {index = n, label = label, name = name, arguments = args,
 | 
| 249 | timeout = mirabelle_timeout, output_dir = output_dir}; | |
| 73847 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 250 | in | 
| 74948 
15ce207f69c8
added support for initialization messages to Mirabelle
 desharna parents: 
74745diff
changeset | 251 | (initialize_action make_action context, context) | 
| 73847 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 252 | end); | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 253 | in | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 254 | (* run actions on all relevant goals *) | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 255 | loaded_theories | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 256 | |> map_index I | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 257 | |> maps make_commands | 
| 74986 
fc664e4fbf6d
added Mirabelle option -r to randomize the goals before selection
 desharna parents: 
74948diff
changeset | 258 | |> (if mirabelle_randomize <= 0 then | 
| 
fc664e4fbf6d
added Mirabelle option -r to randomize the goals before selection
 desharna parents: 
74948diff
changeset | 259 | I | 
| 
fc664e4fbf6d
added Mirabelle option -r to randomize the goals before selection
 desharna parents: 
74948diff
changeset | 260 | else | 
| 
fc664e4fbf6d
added Mirabelle option -r to randomize the goals before selection
 desharna parents: 
74948diff
changeset | 261 | fst o MLCG.shuffle (MLCG.initialize mirabelle_randomize)) | 
| 74077 
b93d8c2ebab0
added automatic uniform stride option to Mirabelle
 desharna parents: 
74069diff
changeset | 262 | |> (fn xs => fold_map (fn x => fn i => ((i, x), i + 1)) xs 0) | 
| 
b93d8c2ebab0
added automatic uniform stride option to Mirabelle
 desharna parents: 
74069diff
changeset | 263 | |> (fn (indexed_commands, commands_length) => | 
| 
b93d8c2ebab0
added automatic uniform stride option to Mirabelle
 desharna parents: 
74069diff
changeset | 264 | let | 
| 
b93d8c2ebab0
added automatic uniform stride option to Mirabelle
 desharna parents: 
74069diff
changeset | 265 | val stride = | 
| 
b93d8c2ebab0
added automatic uniform stride option to Mirabelle
 desharna parents: 
74069diff
changeset | 266 | if mirabelle_stride <= 0 then | 
| 
b93d8c2ebab0
added automatic uniform stride option to Mirabelle
 desharna parents: 
74069diff
changeset | 267 | Integer.max 1 (commands_length div mirabelle_max_calls) | 
| 
b93d8c2ebab0
added automatic uniform stride option to Mirabelle
 desharna parents: 
74069diff
changeset | 268 | else | 
| 
b93d8c2ebab0
added automatic uniform stride option to Mirabelle
 desharna parents: 
74069diff
changeset | 269 | mirabelle_stride | 
| 
b93d8c2ebab0
added automatic uniform stride option to Mirabelle
 desharna parents: 
74069diff
changeset | 270 | in | 
| 
b93d8c2ebab0
added automatic uniform stride option to Mirabelle
 desharna parents: 
74069diff
changeset | 271 | maps (fn (n, command) => | 
| 
b93d8c2ebab0
added automatic uniform stride option to Mirabelle
 desharna parents: 
74069diff
changeset | 272 | let val (m, k) = Integer.div_mod (n + 1) stride in | 
| 
b93d8c2ebab0
added automatic uniform stride option to Mirabelle
 desharna parents: 
74069diff
changeset | 273 | if k = 0 andalso (mirabelle_max_calls <= 0 orelse m <= mirabelle_max_calls) then | 
| 
b93d8c2ebab0
added automatic uniform stride option to Mirabelle
 desharna parents: 
74069diff
changeset | 274 | map (fn context => (context, command)) contexts | 
| 
b93d8c2ebab0
added automatic uniform stride option to Mirabelle
 desharna parents: 
74069diff
changeset | 275 | else | 
| 
b93d8c2ebab0
added automatic uniform stride option to Mirabelle
 desharna parents: 
74069diff
changeset | 276 | [] | 
| 
b93d8c2ebab0
added automatic uniform stride option to Mirabelle
 desharna parents: 
74069diff
changeset | 277 | end) indexed_commands | 
| 73847 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 278 | end) | 
| 74989 | 279 | (* Don't use multithreading for a dry run because of the high per-thread overhead. *) | 
| 280 | |> (if mirabelle_dry_run then map else Par_List.map) (fn ((action, context), command) => | |
| 281 | apply_action (if mirabelle_dry_run then dry_run_action else action) context command); | |
| 73847 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 282 | |
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 283 | (* finalize actions *) | 
| 
58f6b41efe88
refactored Mirabelle to produce output in real time
 desharna parents: 
73822diff
changeset | 284 | List.app (uncurry finalize_action) contexts | 
| 73850 | 285 | end | 
| 73821 
9ead8d9be3ab
clarified hook for Mirabelle: provide all loaded theories at once (for each 'theories' section within the session ROOT);
 wenzelm parents: 
73820diff
changeset | 286 | end); | 
| 32381 | 287 | |
| 288 | ||
| 289 | (* Mirabelle utility functions *) | |
| 290 | ||
| 32472 
7b92a8b8daaf
Mirabelle: actions are responsible for their log messages, output is better readable
 boehmes parents: 
32469diff
changeset | 291 | fun can_apply time tac st = | 
| 32469 | 292 | let | 
| 73691 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 293 |     val {context = ctxt, facts, goal} = Proof.goal st;
 | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 294 | val full_tac = HEADGOAL (Method.insert_tac ctxt facts THEN' tac ctxt); | 
| 32381 | 295 | in | 
| 62519 | 296 | (case try (Timeout.apply time (Seq.pull o full_tac)) goal of | 
| 39451 
8893562a954b
prevent exception when calling "Mirabelle.can_apply" on empty proof sequence;
 blanchet parents: 
39377diff
changeset | 297 | SOME (SOME _) => true | 
| 
8893562a954b
prevent exception when calling "Mirabelle.can_apply" on empty proof sequence;
 blanchet parents: 
39377diff
changeset | 298 | | _ => false) | 
| 73691 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 299 | end; | 
| 32381 | 300 | |
| 301 | local | |
| 302 | ||
| 303 | fun fold_body_thms f = | |
| 304 | let | |
| 64573 | 305 |     fun app n (PBody {thms, ...}) = thms |> fold (fn (i, thm_node) =>
 | 
| 32381 | 306 | fn (x, seen) => | 
| 307 | if Inttab.defined seen i then (x, seen) | |
| 308 | else | |
| 309 | let | |
| 73691 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 310 | val name = Proofterm.thm_node_name thm_node; | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 311 | val prop = Proofterm.thm_node_prop thm_node; | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 312 | val body = Future.join (Proofterm.thm_node_body thm_node); | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 313 | val (x', seen') = | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 314 | app (n + (if name = "" then 0 else 1)) body | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 315 | (x, Inttab.update (i, ()) seen); | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 316 | in (x' |> n = 0 ? f (name, prop, body), seen') end); | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 317 | in fn bodies => fn x => #1 (fold (app 0) bodies (x, Inttab.empty)) end; | 
| 32381 | 318 | |
| 319 | in | |
| 320 | ||
| 60361 | 321 | fun theorems_in_proof_term thy thm = | 
| 32381 | 322 | let | 
| 73691 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 323 | val all_thms = Global_Theory.all_thms_of thy true; | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 324 | fun collect (s, _, _) = if s <> "" then insert (op =) s else I; | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 325 | fun member_of xs (x, y) = if member (op =) xs x then SOME y else NONE; | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 326 | fun resolve_thms names = map_filter (member_of names) all_thms; | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 327 | in resolve_thms (fold_body_thms collect [Thm.proof_body_of thm] []) end; | 
| 32381 | 328 | |
| 73691 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 329 | end; | 
| 32381 | 330 | |
| 73691 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 331 | fun theorems_of_sucessful_proof st = | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 332 | (case try Toplevel.proof_of st of | 
| 32381 | 333 | NONE => [] | 
| 73691 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 334 | | SOME prf => theorems_in_proof_term (Proof.theory_of prf) (#goal (Proof.goal prf))); | 
| 32381 | 335 | |
| 73691 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 336 | fun get_argument arguments (key, default) = | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 337 | the_default default (AList.lookup (op =) arguments key); | 
| 32381 | 338 | |
| 73691 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 339 | fun get_int_argument arguments (key, default) = | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 340 | (case Option.map Int.fromString (AList.lookup (op =) arguments key) of | 
| 32381 | 341 | SOME (SOME i) => i | 
| 342 |   | SOME NONE => error ("bad option: " ^ key)
 | |
| 73691 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 343 | | NONE => default); | 
| 32381 | 344 | |
| 73691 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 345 | fun get_bool_argument arguments (key, default) = | 
| 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 346 | (case Option.map Bool.fromString (AList.lookup (op =) arguments key) of | 
| 73292 
f84a93f1de2f
tuned Mirabelle to parse option check_trivial only once
 desharna parents: 
69597diff
changeset | 347 | SOME (SOME i) => i | 
| 
f84a93f1de2f
tuned Mirabelle to parse option check_trivial only once
 desharna parents: 
69597diff
changeset | 348 |   | SOME NONE => error ("bad option: " ^ key)
 | 
| 73691 
2f9877db82a1
reimplemented Mirabelle as Isabelle/ML presentation hook + Isabelle/Scala tool, but sledgehammer is still inactive;
 wenzelm parents: 
73685diff
changeset | 349 | | NONE => default); | 
| 73292 
f84a93f1de2f
tuned Mirabelle to parse option check_trivial only once
 desharna parents: 
69597diff
changeset | 350 | |
| 42014 
75417ef605ba
simplified various cpu_time clones (!): eliminated odd Exn.capture/Exn.release (no need to "stop" timing);
 wenzelm parents: 
42012diff
changeset | 351 | fun cpu_time f x = | 
| 74991 
d699eb2d26ad
used elapsed time instead of cpu time in Mirabelle because the latter contain cpu time of all threads
 desharna parents: 
74989diff
changeset | 352 | (* CPU time is problematics with multithreading as it refers to the per-process CPU time. *) | 
| 
d699eb2d26ad
used elapsed time instead of cpu time in Mirabelle because the latter contain cpu time of all threads
 desharna parents: 
74989diff
changeset | 353 |   let val ({elapsed, ...}, y) = Timing.timing f x
 | 
| 
d699eb2d26ad
used elapsed time instead of cpu time in Mirabelle because the latter contain cpu time of all threads
 desharna parents: 
74989diff
changeset | 354 | in (y, Time.toMilliseconds elapsed) end; | 
| 32498 
1132c7c13f36
Mirabelle: actions are responsible for handling exceptions,
 boehmes parents: 
32497diff
changeset | 355 | |
| 32381 | 356 | end |