82360
|
1 |
(* Title: HOL/Tools/try0_hol.ML
|
|
2 |
Author: Jasmin Blanchette, LMU Muenchen
|
|
3 |
Author: Martin Desharnais, LMU Muenchen
|
|
4 |
Author: Fabian Huch, TU Muenchen
|
|
5 |
|
|
6 |
General-purpose functions used by Try0.
|
|
7 |
*)
|
|
8 |
|
|
9 |
signature TRY0_UTIL =
|
|
10 |
sig
|
|
11 |
val silence_methods : bool -> Proof.context -> Proof.context
|
|
12 |
val string_of_xref : Try0.xref -> string
|
82361
|
13 |
val apply_raw_named_method : string -> bool -> (Try0.modifier * string) list ->
|
|
14 |
Time.time option -> Try0.tagged_xref list -> Proof.state -> Try0.result option
|
82360
|
15 |
end
|
|
16 |
|
|
17 |
structure Try0_Util : TRY0_UTIL = struct
|
|
18 |
|
|
19 |
(* Makes reconstructor tools as silent as possible. The "set_visible" calls suppresses "Unification
|
|
20 |
bound exceeded" warnings and the like. *)
|
|
21 |
fun silence_methods debug =
|
|
22 |
Config.put Metis_Tactic.verbose debug
|
|
23 |
#> not debug ? (fn ctxt =>
|
|
24 |
ctxt
|
|
25 |
|> Simplifier_Trace.disable
|
|
26 |
|> Context_Position.set_visible false
|
|
27 |
|> Config.put Unify.unify_trace false
|
|
28 |
|> Config.put Argo_Tactic.trace "none")
|
|
29 |
|
|
30 |
fun string_of_xref ((xref, args) : Try0.xref) =
|
|
31 |
(case xref of
|
|
32 |
Facts.Fact literal => literal |> Symbol_Pos.explode0 |> Symbol_Pos.implode |> cartouche
|
|
33 |
| _ =>
|
|
34 |
Facts.string_of_ref xref) ^ implode
|
|
35 |
(map (enclose "[" "]" o Pretty.unformatted_string_of o Token.pretty_src \<^context>) args)
|
|
36 |
|
82361
|
37 |
local
|
|
38 |
|
|
39 |
fun add_attr_text (tagged : Try0.tagged_xref list) (tag, src) s =
|
|
40 |
let
|
|
41 |
val fs = tagged
|
|
42 |
|> map_filter (fn (xref, tags) =>
|
|
43 |
if member (op =) tags tag then
|
|
44 |
SOME (string_of_xref xref)
|
|
45 |
else
|
|
46 |
NONE)
|
|
47 |
in if null fs then s else s ^ " " ^ (if src = "" then "" else src ^ ": ") ^ implode_space fs end
|
|
48 |
|
|
49 |
fun attrs_text tags (tagged : Try0.tagged_xref list) =
|
|
50 |
"" |> fold (add_attr_text tagged) tags
|
|
51 |
|
|
52 |
fun parse_method ctxt s =
|
|
53 |
enclose "(" ")" s
|
|
54 |
|> Token.explode (Thy_Header.get_keywords' ctxt) Position.start
|
|
55 |
|> filter Token.is_proper
|
|
56 |
|> Scan.read Token.stopper Method.parse
|
|
57 |
|> (fn SOME (Method.Source src, _) => src | _ => raise Fail "expected Source")
|
|
58 |
|
|
59 |
fun run_tac timeout_opt tac st =
|
|
60 |
let val with_timeout =
|
|
61 |
(case timeout_opt of SOME timeout => Timeout.apply_physical timeout | NONE => I)
|
|
62 |
in with_timeout (Seq.pull o tac) st |> Option.map fst end
|
|
63 |
|
|
64 |
val num_goals = Thm.nprems_of o #goal o Proof.goal
|
|
65 |
|
|
66 |
fun apply_recursive recurse elapsed0 timeout_opt apply st =
|
|
67 |
(case Timing.timing (Option.join o try (run_tac timeout_opt apply)) st of
|
|
68 |
({elapsed, ...}, SOME st') =>
|
|
69 |
if recurse andalso num_goals st' > 0 andalso num_goals st' < num_goals st then
|
|
70 |
let val timeout_opt1 = (Option.map (fn timeout => timeout - elapsed) timeout_opt)
|
|
71 |
in apply_recursive recurse (elapsed0 + elapsed) timeout_opt1 apply st' end
|
|
72 |
else (elapsed0 + elapsed, st')
|
|
73 |
|_ => (elapsed0, st))
|
|
74 |
|
|
75 |
in
|
|
76 |
|
|
77 |
fun apply_raw_named_method (name : string) all_goals attrs timeout_opt
|
|
78 |
(tagged : Try0.tagged_xref list) (st : Proof.state) : Try0.result option =
|
|
79 |
let
|
|
80 |
val st = Proof.map_contexts (silence_methods false) st
|
|
81 |
val unused =
|
|
82 |
tagged
|
|
83 |
|> filter
|
|
84 |
(fn (_, tags) => not (null tags) andalso null (inter (op =) tags (attrs |> map fst)))
|
|
85 |
|> map fst
|
|
86 |
|
|
87 |
val attrs = attrs_text attrs tagged
|
|
88 |
|
|
89 |
val ctxt = Proof.context_of st
|
|
90 |
|
|
91 |
val text =
|
|
92 |
name ^ attrs
|
|
93 |
|> parse_method ctxt
|
|
94 |
|> Method.method_cmd ctxt
|
|
95 |
|> Method.Basic
|
|
96 |
|> (fn m => Method.Combinator (Method.no_combinator_info, Method.Select_Goals 1, [m]))
|
|
97 |
|
|
98 |
val apply =
|
|
99 |
Proof.using [Attrib.eval_thms ctxt unused |> map (rpair [] o single)]
|
|
100 |
#> Proof.refine text #> Seq.filter_results
|
|
101 |
val num_before = num_goals st
|
|
102 |
val multiple_goals = all_goals andalso num_before > 1
|
|
103 |
val (time, st') = apply_recursive multiple_goals Time.zeroTime timeout_opt apply st
|
|
104 |
val num_after = num_goals st'
|
|
105 |
val select = "[" ^ string_of_int (num_before - num_after) ^ "]"
|
|
106 |
val unused = implode_space (unused |> map string_of_xref)
|
|
107 |
val command =
|
|
108 |
(if unused <> "" then "using " ^ unused ^ " " else "") ^
|
|
109 |
(if num_after = 0 then "by " else "apply ") ^
|
|
110 |
(name ^ attrs |> attrs <> "" ? enclose "(" ")") ^
|
|
111 |
(if multiple_goals andalso num_after > 0 then select else "")
|
|
112 |
in
|
|
113 |
if num_before > num_after then
|
|
114 |
SOME {name = name, command = command, time = time, state = st'}
|
|
115 |
else NONE
|
|
116 |
end
|
|
117 |
|
|
118 |
end
|
|
119 |
|
82360
|
120 |
end |