-
Notifications
You must be signed in to change notification settings - Fork 0
/
Apply_Trace.thy
319 lines (249 loc) · 9.91 KB
/
Apply_Trace.thy
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
274
275
276
277
278
279
280
281
282
283
284
285
286
287
288
289
290
291
292
293
294
295
296
297
298
299
300
301
302
303
304
305
306
307
308
309
310
311
312
313
314
315
316
317
318
(*
* Copyright 2014, NICTA
*
* This software may be distributed and modified according to the terms of
* the BSD 2-Clause license. Note that NO WARRANTY is provided.
* See "LICENSE_BSD2.txt" for details.
*
* @TAG(NICTA_BSD)
*)
(* Backend for tracing apply statements. Useful for doing proof step dependency analysis.
* Provides an alternate refinement function which takes an additional stateful journaling operation. *)
theory Apply_Trace
imports Main
begin
ML {*
signature APPLY_TRACE =
sig
val apply_results :
{silent_fail : bool} ->
(Proof.context -> thm -> ((string * int option) * term) list -> unit) ->
Method.text_range -> Proof.state -> Proof.state Seq.result Seq.seq
(* Lower level interface. *)
val can_clear : theory -> bool
val clear_deps : thm -> thm
val join_deps : thm -> thm -> thm
val used_facts : Proof.context -> thm -> ((string * int option) * term) list
val pretty_deps: bool -> (string * Position.T) option -> Proof.context -> thm ->
((string * int option) * term) list -> Pretty.T
end
structure Apply_Trace : APPLY_TRACE =
struct
(*TODO: Add more robust oracle without hyp clearing *)
fun thm_to_cterm keep_hyps thm =
let
val thy = Thm.theory_of_thm thm
val pairs = Thm.tpairs_of thm
val ceqs = map (Thm.global_cterm_of thy o Logic.mk_equals) pairs
val hyps = Thm.chyps_of thm
val prop = Thm.cprop_of thm
val thm' = if keep_hyps then Drule.list_implies (hyps,prop) else prop
in
Drule.list_implies (ceqs,thm') end
val (_, clear_thm_deps') =
Context.>>> (Context.map_theory_result (Thm.add_oracle (Binding.name "count_cheat", thm_to_cterm false)));
fun clear_deps thm =
let
val thm' = try clear_thm_deps' thm
|> Option.map (fold (fn _ => fn t => (@{thm Pure.reflexive} RS t)) (Thm.tpairs_of thm))
in case thm' of SOME thm' => thm' | NONE => error "Can't clear deps here" end
fun can_clear thy = Context.subthy(@{theory},thy)
fun join_deps pre_thm post_thm =
let
val pre_thm' = Thm.flexflex_rule NONE pre_thm |> Seq.hd
|> Thm.adjust_maxidx_thm (Thm.maxidx_of post_thm + 1)
in
Conjunction.intr pre_thm' post_thm |> Conjunction.elim |> snd
end
fun get_ref_from_nm' nm =
let
val exploded = space_explode "_" nm;
val base = List.take (exploded, (length exploded) - 1) |> space_implode "_"
val idx = List.last exploded |> Int.fromString;
in if is_some idx andalso base <> "" then SOME (base, the idx) else NONE end
fun get_ref_from_nm nm = Option.join (try get_ref_from_nm' nm);
fun maybe_nth l = try (curry List.nth l)
fun fact_from_derivation ctxt xnm =
let
val facts = Proof_Context.facts_of ctxt;
(* TODO: Check that exported local fact is equivalent to external one *)
val idx_result =
let
val (name', idx) = get_ref_from_nm xnm |> the;
val entry = try (Facts.retrieve (Context.Proof ctxt) facts) (name', Position.none) |> the;
val thm = maybe_nth (#thms entry) (idx - 1) |> the;
in SOME (#name entry, thm) end handle Option => NONE;
fun non_idx_result () =
let
val entry = try (Facts.retrieve (Context.Proof ctxt) facts) (xnm, Position.none) |> the;
val thm = try the_single (#thms entry) |> the;
in SOME (#name entry, thm) end handle Option => NONE;
in
case idx_result of
SOME thm => SOME thm
| NONE => non_idx_result ()
end
fun most_local_fact_of ctxt xnm =
let
val local_name = try (fn xnm => Long_Name.explode xnm |> tl |> tl |> Long_Name.implode) xnm |> the;
in SOME (fact_from_derivation ctxt local_name |> the) end handle Option =>
fact_from_derivation ctxt xnm;
fun thms_of (PBody {thms,...}) = thms
fun proof_body_descend' f get_fact (ident,(nm,_ , body)) deptab =
(if not (f nm) then
(Inttab.update_new (ident, SOME (nm, get_fact nm |> the)) deptab handle Inttab.DUP _ => deptab)
else raise Option) handle Option =>
((fold (proof_body_descend' f get_fact) (thms_of (Future.join body))
(Inttab.update_new (ident, NONE) deptab)) handle Inttab.DUP _ => deptab)
fun used_facts' f get_fact thm =
let
val body = thms_of (Thm.proof_body_of thm);
in fold (proof_body_descend' f get_fact) body Inttab.empty end
fun used_pbody_facts ctxt thm =
let
val nm = Thm.get_name_hint thm;
val get_fact = most_local_fact_of ctxt;
in
used_facts' (fn nm' => nm' = "" orelse nm' = nm) get_fact thm
|> Inttab.dest |> map_filter snd |> map snd |> map (apsnd (Thm.prop_of))
end
fun raw_primitive_text f = Method.Basic (fn _ => ((K (fn (ctxt, thm) => Seq.make_results (Seq.single (ctxt, f thm))))))
(*Find local facts from new hyps*)
fun used_local_facts ctxt thm =
let
val hyps = Thm.hyps_of thm
val facts = Proof_Context.facts_of ctxt |> Facts.dest_static true []
fun match_hyp hyp =
let
fun get (nm,thms) =
case (get_index (fn t => if (Thm.prop_of t) aconv hyp then SOME hyp else NONE) thms)
of SOME t => SOME (nm,t)
| NONE => NONE
in
get_first get facts
end
in
map_filter match_hyp hyps end
fun used_facts ctxt thm =
let
val used_from_pbody = used_pbody_facts ctxt thm |> map (fn (nm,t) => ((nm,NONE),t))
val used_from_hyps = used_local_facts ctxt thm |> map (fn (nm,(i,t)) => ((nm,SOME i),t))
in
(used_from_hyps @ used_from_pbody)
end
(* Perform refinement step, and run the given stateful function
against computed dependencies afterwards. *)
fun refine args f text state =
let
val ctxt = Proof.context_of state
val thm = Proof.simple_goal state |> #goal
fun save_deps deps = f ctxt thm deps
in
if (can_clear (Proof.theory_of state)) then
Proof.refine (Method.Combinator (Method.no_combinator_info,Method.Then, [raw_primitive_text (clear_deps),text,
raw_primitive_text (fn thm' => (save_deps (used_facts ctxt thm');join_deps thm thm'))])) state
else
(if (#silent_fail args) then (save_deps [];Proof.refine text state) else error "Apply_Trace theory must be imported to trace applies")
end
(* Boilerplate from Proof.ML *)
fun method_error kind pos state =
Seq.single (Proof_Display.method_error kind pos (Proof.raw_goal state));
fun apply args f text = Proof.assert_backward #> refine args f text #>
Seq.maps_results (Proof.apply ((raw_primitive_text I),(Position.none, Position.none)));
fun apply_results args f (text, range) =
Seq.APPEND (apply args f text, method_error "" (Position.range_position range));
structure Filter_Thms = Named_Thms
(
val name = @{binding no_trace}
val description = "thms to be ignored from tracing"
)
datatype adjusted_name =
FoundName of ((string * int option) * thm)
| UnknownName of (string * term)
(* Parse the index of a theorem name in the form "x_1". *)
fun parse_thm_index name =
case (String.tokens (fn c => c = #"_") name |> rev) of
(possible_index::xs) =>
(case Lexicon.read_nat possible_index of
SOME n => (space_implode "_" (rev xs), SOME (n - 1))
| NONE => (name, NONE))
| _ => (name, NONE)
(*
* Names stored in proof bodies may have the form "x_1" which can either
* mean "x(1)" or "x_1". Attempt to determine the correct name for the
* given theorem. If we can't find the correct theorem, or it is
* ambiguous, return the original name.
*)
fun adjust_thm_name ctxt (name,index) term =
let
val possible_names = case index of NONE => distinct (op =) [(name, NONE), parse_thm_index name]
| SOME i => [(name,SOME i)]
fun match (n, i) =
let
val idx = the_default 0 i
val thms = Proof_Context.get_fact ctxt (Facts.named n) handle ERROR _ => []
in
if idx >= 0 andalso length thms > idx then
if length thms > 1 then
SOME ((n, i), nth thms idx)
else
SOME ((n,NONE), hd thms)
else
NONE
end
in
case map_filter match possible_names of
[x] => FoundName x
| _ => UnknownName (name, term)
end
(* Render the given fact. *)
fun pretty_fact only_names ctxt (FoundName ((name, idx), thm)) =
Pretty.block
([Pretty.mark_str (Facts.markup_extern ctxt (Proof_Context.facts_of ctxt) name),
case idx of
SOME n => Pretty.str ("(" ^ string_of_int (n + 1) ^ ")")
| NONE => Pretty.str ""] @
(if only_names then []
else [Pretty.str ":",Pretty.brk 1, Thm.pretty_thm ctxt thm]))
| pretty_fact _ ctxt (UnknownName (name, prop)) =
Pretty.block
[Pretty.str name, Pretty.str "(?) :", Pretty.brk 1,
Syntax.unparse_term ctxt prop]
fun fact_ref_to_name ((Facts.Named ((nm,_), (SOME [Facts.Single i]))),thm) = FoundName ((nm,SOME i),thm)
| fact_ref_to_name ((Facts.Named ((nm,_), (NONE))),thm) = FoundName ((nm,NONE),thm)
| fact_ref_to_name (_,thm) = UnknownName ("",Thm.prop_of thm)
(* Print out the found dependencies. *)
fun pretty_deps only_names query ctxt thm deps =
let
(* Remove duplicates. *)
val deps = sort_distinct (prod_ord (prod_ord string_ord (option_ord int_ord)) Term_Ord.term_ord) deps
(* Fetch canonical names and theorems. *)
val deps = map (fn (ident, term) => adjust_thm_name ctxt ident term) deps
(* Remove "boring" theorems. *)
val deps = subtract (fn (a, FoundName (_, thm)) => Thm.eq_thm (thm, a)
| _ => false) (Filter_Thms.get ctxt) deps
val deps = case query of SOME (raw_query,pos) =>
let
val pos' = perhaps (try (Position.advance_offset 1)) pos;
val q = Find_Theorems.read_query pos' raw_query;
val results = Find_Theorems.find_theorems_cmd ctxt (SOME thm) (SOME 1000000000) false q
|> snd
|> map fact_ref_to_name;
(* Only consider theorems from our query. *)
val deps = inter (fn (FoundName (nmidx,_), FoundName (nmidx',_)) => nmidx = nmidx'
| _ => false) results deps
in deps end
| _ => deps
in
if only_names then
Pretty.block
(Pretty.separate "" (map ((pretty_fact only_names) ctxt) deps))
else
(* Pretty-print resulting theorems. *)
Pretty.big_list "used theorems:"
(map (Pretty.item o single o (pretty_fact only_names) ctxt) deps)
end
val _ = Context.>> (Context.map_theory Filter_Thms.setup)
end
*}
end