| author | wenzelm |
| Sat, 01 Feb 2014 21:43:23 +0100 | |
| changeset 55240 | efc4c0e0e14a |
| parent 55223 | 3c593bad6b31 |
| child 55244 | 12e1a5d8ee48 |
| permissions | -rw-r--r-- |
|
55202
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
blanchet
parents:
55194
diff
changeset
|
1 |
(* Title: HOL/Tools/Sledgehammer/sledgehammer_isar_proof.ML |
| 50263 | 2 |
Author: Jasmin Blanchette, TU Muenchen |
3 |
Author: Steffen Juilf Smolka, TU Muenchen |
|
4 |
||
5 |
Basic data structures for representing and basic methods |
|
6 |
for dealing with Isar proof texts. |
|
7 |
*) |
|
8 |
||
|
55202
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
blanchet
parents:
55194
diff
changeset
|
9 |
signature SLEDGEHAMMER_ISAR_PROOF = |
| 50259 | 10 |
sig |
| 51239 | 11 |
type label = string * int |
|
54816
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
blanchet
parents:
54767
diff
changeset
|
12 |
type facts = label list * string list (* local and global facts *) |
| 50268 | 13 |
|
| 51178 | 14 |
datatype isar_qualifier = Show | Then |
| 50268 | 15 |
|
| 55218 | 16 |
datatype proof_method = |
17 |
Metis_Method | Simp_Method | Simp_Size_Method | Auto_Method | Fastforce_Method | Force_Method | |
|
| 55219 | 18 |
Arith_Method | Blast_Method | Meson_Method | Algebra_Method |
| 55218 | 19 |
|
| 52454 | 20 |
datatype isar_proof = |
| 54700 | 21 |
Proof of (string * typ) list * (label * term) list * isar_step list |
|
51179
0d5f8812856f
split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents:
51178
diff
changeset
|
22 |
and isar_step = |
| 50268 | 23 |
Let of term * term | |
| 54700 | 24 |
Prove of isar_qualifier list * (string * typ) list * label * term * isar_proof list * |
|
54767
81290fe85890
generalize method list further to list of list (clustering preferred methods together)
blanchet
parents:
54766
diff
changeset
|
25 |
(facts * proof_method list list) |
|
52592
8a25b17e3d79
optimize isar-proofs by trying different proof methods
smolkas
parents:
52590
diff
changeset
|
26 |
|
|
51179
0d5f8812856f
split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents:
51178
diff
changeset
|
27 |
val no_label : label |
|
0d5f8812856f
split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents:
51178
diff
changeset
|
28 |
val no_facts : facts |
| 50268 | 29 |
|
|
52592
8a25b17e3d79
optimize isar-proofs by trying different proof methods
smolkas
parents:
52590
diff
changeset
|
30 |
val label_ord : label * label -> order |
|
51998
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents:
51239
diff
changeset
|
31 |
val string_of_label : label -> string |
|
52592
8a25b17e3d79
optimize isar-proofs by trying different proof methods
smolkas
parents:
52590
diff
changeset
|
32 |
|
| 55222 | 33 |
val string_of_proof_method : proof_method -> string |
34 |
||
|
51179
0d5f8812856f
split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents:
51178
diff
changeset
|
35 |
val steps_of_proof : isar_proof -> isar_step list |
|
0d5f8812856f
split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents:
51178
diff
changeset
|
36 |
|
|
55223
3c593bad6b31
generalized preplaying infrastructure to store various results for various methods
blanchet
parents:
55222
diff
changeset
|
37 |
val label_of_isar_step : isar_step -> label option |
|
3c593bad6b31
generalized preplaying infrastructure to store various results for various methods
blanchet
parents:
55222
diff
changeset
|
38 |
val byline_of_isar_step : isar_step -> (facts * proof_method list list) option |
|
51179
0d5f8812856f
split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents:
51178
diff
changeset
|
39 |
|
| 54765 | 40 |
val fold_isar_steps : (isar_step -> 'a -> 'a) -> isar_step list -> 'a -> 'a |
| 55212 | 41 |
val map_isar_steps : (isar_step -> isar_step) -> isar_proof -> isar_proof |
42 |
val add_isar_steps : isar_step list -> int -> int |
|
| 52454 | 43 |
|
| 55212 | 44 |
structure Canonical_Label_Tab : TABLE |
|
52556
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents:
52454
diff
changeset
|
45 |
|
|
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents:
52454
diff
changeset
|
46 |
val canonical_label_ord : (label * label) -> order |
| 55220 | 47 |
|
48 |
val chain_isar_proof : isar_proof -> isar_proof |
|
49 |
val kill_useless_labels_in_isar_proof : isar_proof -> isar_proof |
|
| 55213 | 50 |
val relabel_isar_proof_canonically : isar_proof -> isar_proof |
| 55220 | 51 |
val relabel_isar_proof_finally : isar_proof -> isar_proof |
|
52556
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents:
52454
diff
changeset
|
52 |
|
| 55222 | 53 |
val string_of_isar_proof : Proof.context -> string -> string -> int -> int -> |
54 |
(label -> proof_method list list -> string) -> isar_proof -> string |
|
| 54504 | 55 |
end; |
| 50259 | 56 |
|
|
55202
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
blanchet
parents:
55194
diff
changeset
|
57 |
structure Sledgehammer_Isar_Proof : SLEDGEHAMMER_ISAR_PROOF = |
| 50259 | 58 |
struct |
59 |
||
| 55211 | 60 |
open ATP_Util |
61 |
open ATP_Proof |
|
62 |
open ATP_Problem_Generate |
|
63 |
open ATP_Proof_Reconstruct |
|
64 |
open Sledgehammer_Util |
|
65 |
open Sledgehammer_Reconstructor |
|
66 |
open Sledgehammer_Isar_Annotate |
|
67 |
||
| 50259 | 68 |
type label = string * int |
| 54534 | 69 |
type facts = label list * string list (* local and global facts *) |
| 50259 | 70 |
|
| 51178 | 71 |
datatype isar_qualifier = Show | Then |
| 50259 | 72 |
|
| 55218 | 73 |
datatype proof_method = |
74 |
Metis_Method | Simp_Method | Simp_Size_Method | Auto_Method | Fastforce_Method | Force_Method | |
|
| 55219 | 75 |
Arith_Method | Blast_Method | Meson_Method | Algebra_Method |
| 55218 | 76 |
|
| 52454 | 77 |
datatype isar_proof = |
| 54700 | 78 |
Proof of (string * typ) list * (label * term) list * isar_step list |
|
51179
0d5f8812856f
split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents:
51178
diff
changeset
|
79 |
and isar_step = |
| 50259 | 80 |
Let of term * term | |
| 54700 | 81 |
Prove of isar_qualifier list * (string * typ) list * label * term * isar_proof list * |
|
54767
81290fe85890
generalize method list further to list of list (clustering preferred methods together)
blanchet
parents:
54766
diff
changeset
|
82 |
(facts * proof_method list list) |
|
52592
8a25b17e3d79
optimize isar-proofs by trying different proof methods
smolkas
parents:
52590
diff
changeset
|
83 |
|
|
51179
0d5f8812856f
split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents:
51178
diff
changeset
|
84 |
val no_label = ("", ~1)
|
|
0d5f8812856f
split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents:
51178
diff
changeset
|
85 |
val no_facts = ([],[]) |
| 50259 | 86 |
|
|
52592
8a25b17e3d79
optimize isar-proofs by trying different proof methods
smolkas
parents:
52590
diff
changeset
|
87 |
val label_ord = pairself swap #> prod_ord int_ord fast_string_ord |
|
52556
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents:
52454
diff
changeset
|
88 |
|
|
51998
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents:
51239
diff
changeset
|
89 |
fun string_of_label (s, num) = s ^ string_of_int num |
| 50259 | 90 |
|
| 55222 | 91 |
fun string_of_proof_method meth = |
| 55218 | 92 |
(case meth of |
93 |
Metis_Method => "metis" |
|
94 |
| Simp_Method => "simp" |
|
95 |
| Simp_Size_Method => "(simp add: size_ne_size_imp_ne)" |
|
96 |
| Auto_Method => "auto" |
|
97 |
| Fastforce_Method => "fastforce" |
|
98 |
| Force_Method => "force" |
|
99 |
| Arith_Method => "arith" |
|
100 |
| Blast_Method => "blast" |
|
| 55219 | 101 |
| Meson_Method => "meson" |
102 |
| Algebra_Method => "algebra") |
|
| 55218 | 103 |
|
|
51179
0d5f8812856f
split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents:
51178
diff
changeset
|
104 |
fun steps_of_proof (Proof (_, _, steps)) = steps |
|
0d5f8812856f
split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents:
51178
diff
changeset
|
105 |
|
|
55223
3c593bad6b31
generalized preplaying infrastructure to store various results for various methods
blanchet
parents:
55222
diff
changeset
|
106 |
fun label_of_isar_step (Prove (_, _, l, _, _, _)) = SOME l |
|
3c593bad6b31
generalized preplaying infrastructure to store various results for various methods
blanchet
parents:
55222
diff
changeset
|
107 |
| label_of_isar_step _ = NONE |
| 51178 | 108 |
|
|
55223
3c593bad6b31
generalized preplaying infrastructure to store various results for various methods
blanchet
parents:
55222
diff
changeset
|
109 |
fun subproofs_of_isar_step (Prove (_, _, _, _, subproofs, _)) = SOME subproofs |
|
3c593bad6b31
generalized preplaying infrastructure to store various results for various methods
blanchet
parents:
55222
diff
changeset
|
110 |
| subproofs_of_isar_step _ = NONE |
| 52454 | 111 |
|
|
55223
3c593bad6b31
generalized preplaying infrastructure to store various results for various methods
blanchet
parents:
55222
diff
changeset
|
112 |
fun byline_of_isar_step (Prove (_, _, _, _, _, byline)) = SOME byline |
|
3c593bad6b31
generalized preplaying infrastructure to store various results for various methods
blanchet
parents:
55222
diff
changeset
|
113 |
| byline_of_isar_step _ = NONE |
|
51179
0d5f8812856f
split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents:
51178
diff
changeset
|
114 |
|
| 55212 | 115 |
fun fold_isar_step f step = |
|
55223
3c593bad6b31
generalized preplaying infrastructure to store various results for various methods
blanchet
parents:
55222
diff
changeset
|
116 |
fold (steps_of_proof #> fold_isar_steps f) (these (subproofs_of_isar_step step)) #> f step |
| 55212 | 117 |
and fold_isar_steps f = fold (fold_isar_step f) |
| 52454 | 118 |
|
| 54765 | 119 |
fun map_isar_steps f = |
|
52592
8a25b17e3d79
optimize isar-proofs by trying different proof methods
smolkas
parents:
52590
diff
changeset
|
120 |
let |
| 55212 | 121 |
fun map_proof (Proof (fix, assms, steps)) = Proof (fix, assms, map map_step steps) |
122 |
and map_step (step as Let _) = f step |
|
123 |
| map_step (Prove (qs, xs, l, t, subproofs, by)) = |
|
124 |
f (Prove (qs, xs, l, t, map map_proof subproofs, by)) |
|
125 |
in map_proof end |
|
|
52592
8a25b17e3d79
optimize isar-proofs by trying different proof methods
smolkas
parents:
52590
diff
changeset
|
126 |
|
| 55212 | 127 |
val add_isar_steps = fold_isar_steps (fn Prove _ => Integer.add 1 | _ => I) |
|
52556
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents:
52454
diff
changeset
|
128 |
|
| 55211 | 129 |
(* canonical proof labels: 1, 2, 3, ... in post traversal order *) |
| 52557 | 130 |
fun canonical_label_ord (((_, i1), (_, i2)) : label * label) = int_ord (i1, i2) |
|
52556
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents:
52454
diff
changeset
|
131 |
|
| 55212 | 132 |
structure Canonical_Label_Tab = Table( |
|
52556
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents:
52454
diff
changeset
|
133 |
type key = label |
|
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents:
52454
diff
changeset
|
134 |
val ord = canonical_label_ord) |
|
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents:
52454
diff
changeset
|
135 |
|
| 55220 | 136 |
fun chain_qs_lfs NONE lfs = ([], lfs) |
137 |
| chain_qs_lfs (SOME l0) lfs = |
|
138 |
if member (op =) lfs l0 then ([Then], lfs |> remove (op =) l0) else ([], lfs) |
|
139 |
fun chain_isar_step lbl (Prove (qs, xs, l, t, subproofs, ((lfs, gfs), methss))) = |
|
140 |
let val (qs', lfs) = chain_qs_lfs lbl lfs in |
|
141 |
Prove (qs' @ qs, xs, l, t, map chain_isar_proof subproofs, ((lfs, gfs), methss)) |
|
142 |
end |
|
143 |
| chain_isar_step _ step = step |
|
144 |
and chain_isar_steps _ [] = [] |
|
145 |
| chain_isar_steps (prev as SOME _) (i :: is) = |
|
|
55223
3c593bad6b31
generalized preplaying infrastructure to store various results for various methods
blanchet
parents:
55222
diff
changeset
|
146 |
chain_isar_step prev i :: chain_isar_steps (label_of_isar_step i) is |
|
3c593bad6b31
generalized preplaying infrastructure to store various results for various methods
blanchet
parents:
55222
diff
changeset
|
147 |
| chain_isar_steps _ (i :: is) = i :: chain_isar_steps (label_of_isar_step i) is |
| 55220 | 148 |
and chain_isar_proof (Proof (fix, assms, steps)) = |
149 |
Proof (fix, assms, chain_isar_steps (try (List.last #> fst) assms) steps) |
|
150 |
||
151 |
fun kill_useless_labels_in_isar_proof proof = |
|
152 |
let |
|
153 |
val used_ls = |
|
|
55223
3c593bad6b31
generalized preplaying infrastructure to store various results for various methods
blanchet
parents:
55222
diff
changeset
|
154 |
fold_isar_steps (byline_of_isar_step #> (fn SOME ((ls, _), _) => union (op =) ls | _ => I)) |
| 55220 | 155 |
(steps_of_proof proof) [] |
156 |
||
157 |
fun kill_label l = if member (op =) used_ls l then l else no_label |
|
158 |
||
159 |
fun kill_step (Prove (qs, xs, l, t, subproofs, by)) = |
|
160 |
Prove (qs, xs, kill_label l, t, map kill_proof subproofs, by) |
|
161 |
| kill_step step = step |
|
162 |
and kill_proof (Proof (fix, assms, steps)) = |
|
163 |
Proof (fix, map (apfst kill_label) assms, map kill_step steps) |
|
164 |
in |
|
165 |
kill_proof proof |
|
166 |
end |
|
167 |
||
| 55213 | 168 |
fun relabel_isar_proof_canonically proof = |
|
52556
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents:
52454
diff
changeset
|
169 |
let |
|
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents:
52454
diff
changeset
|
170 |
fun next_label l (next, subst) = |
| 54534 | 171 |
let val l' = ("", next) in (l', (next + 1, (l, l') :: subst)) end
|
|
52556
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents:
52454
diff
changeset
|
172 |
|
| 54700 | 173 |
fun do_byline by (_, subst) = apfst (apfst (map (AList.lookup (op =) subst #> the))) by |
|
52556
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents:
52454
diff
changeset
|
174 |
handle Option.Option => |
| 55213 | 175 |
raise Fail "Sledgehammer_Isar_Proof: relabel_isar_proof_canonically" |
|
52556
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents:
52454
diff
changeset
|
176 |
|
|
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents:
52454
diff
changeset
|
177 |
fun do_assm (l, t) state = |
|
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents:
52454
diff
changeset
|
178 |
let |
|
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents:
52454
diff
changeset
|
179 |
val (l, state) = next_label l state |
|
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents:
52454
diff
changeset
|
180 |
in ((l, t), state) end |
|
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents:
52454
diff
changeset
|
181 |
|
| 54700 | 182 |
fun do_proof (Proof (fix, assms, steps)) state = |
|
52556
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents:
52454
diff
changeset
|
183 |
let |
|
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents:
52454
diff
changeset
|
184 |
val (assms, state) = fold_map do_assm assms state |
|
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents:
52454
diff
changeset
|
185 |
val (steps, state) = fold_map do_step steps state |
|
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents:
52454
diff
changeset
|
186 |
in |
| 54700 | 187 |
(Proof (fix, assms, steps), state) |
|
52556
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents:
52454
diff
changeset
|
188 |
end |
|
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents:
52454
diff
changeset
|
189 |
|
|
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents:
52454
diff
changeset
|
190 |
and do_step (step as Let _) state = (step, state) |
|
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents:
52454
diff
changeset
|
191 |
| do_step (Prove (qs, fix, l, t, subproofs, by)) state= |
| 54534 | 192 |
let |
193 |
val by = do_byline by state |
|
194 |
val (subproofs, state) = fold_map do_proof subproofs state |
|
195 |
val (l, state) = next_label l state |
|
196 |
in |
|
197 |
(Prove (qs, fix, l, t, subproofs, by), state) |
|
198 |
end |
|
|
52556
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents:
52454
diff
changeset
|
199 |
in |
|
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents:
52454
diff
changeset
|
200 |
fst (do_proof proof (0, [])) |
|
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents:
52454
diff
changeset
|
201 |
end |
|
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents:
52454
diff
changeset
|
202 |
|
| 55220 | 203 |
val assume_prefix = "a" |
204 |
val have_prefix = "f" |
|
205 |
||
206 |
val relabel_isar_proof_finally = |
|
207 |
let |
|
208 |
fun fresh_label depth prefix (accum as (l, subst, next)) = |
|
209 |
if l = no_label then |
|
210 |
accum |
|
211 |
else |
|
212 |
let val l' = (replicate_string (depth + 1) prefix, next) in |
|
213 |
(l', (l, l') :: subst, next + 1) |
|
214 |
end |
|
215 |
||
216 |
fun relabel_facts subst = apfst (maps (the_list o AList.lookup (op =) subst)) |
|
217 |
||
218 |
fun relabel_assm depth (l, t) (subst, next) = |
|
219 |
let val (l, subst, next) = (l, subst, next) |> fresh_label depth assume_prefix in |
|
220 |
((l, t), (subst, next)) |
|
221 |
end |
|
222 |
||
223 |
fun relabel_assms subst depth assms = fold_map (relabel_assm depth) assms (subst, 1) ||> fst |
|
224 |
||
225 |
fun relabel_steps _ _ _ [] = [] |
|
226 |
| relabel_steps subst depth next (Prove (qs, xs, l, t, sub, by) :: steps) = |
|
227 |
let |
|
228 |
val (l, subst, next) = (l, subst, next) |> fresh_label depth have_prefix |
|
229 |
val sub = relabel_proofs subst depth sub |
|
230 |
val by = apfst (relabel_facts subst) by |
|
231 |
in |
|
232 |
Prove (qs, xs, l, t, sub, by) :: relabel_steps subst depth next steps |
|
233 |
end |
|
234 |
| relabel_steps subst depth next (step :: steps) = |
|
235 |
step :: relabel_steps subst depth next steps |
|
236 |
and relabel_proof subst depth (Proof (fix, assms, steps)) = |
|
237 |
let val (assms, subst) = relabel_assms subst depth assms in |
|
238 |
Proof (fix, assms, relabel_steps subst depth 1 steps) |
|
239 |
end |
|
240 |
and relabel_proofs subst depth = map (relabel_proof subst (depth + 1)) |
|
241 |
in |
|
242 |
relabel_proof [] 0 |
|
243 |
end |
|
244 |
||
| 55211 | 245 |
val indent_size = 2 |
246 |
||
| 55222 | 247 |
fun string_of_isar_proof ctxt type_enc lam_trans i n comment_of proof = |
| 55211 | 248 |
let |
249 |
(* Make sure only type constraints inserted by the type annotation code are printed. *) |
|
250 |
val ctxt = |
|
251 |
ctxt |> Config.put show_markup false |
|
252 |
|> Config.put Printer.show_type_emphasis false |
|
253 |
|> Config.put show_types false |
|
254 |
|> Config.put show_sorts false |
|
255 |
|> Config.put show_consts false |
|
256 |
||
257 |
val register_fixes = map Free #> fold Variable.auto_fixes |
|
258 |
||
| 55216 | 259 |
fun add_str s' = apfst (suffix s') |
| 55211 | 260 |
|
261 |
fun of_indent ind = replicate_string (ind * indent_size) " " |
|
262 |
fun of_moreover ind = of_indent ind ^ "moreover\n" |
|
263 |
fun of_label l = if l = no_label then "" else string_of_label l ^ ": " |
|
264 |
||
265 |
fun of_obtain qs nr = |
|
266 |
(if nr > 1 orelse (nr = 1 andalso member (op =) qs Then) then "ultimately " |
|
267 |
else if nr = 1 orelse member (op =) qs Then then "then " |
|
268 |
else "") ^ "obtain" |
|
269 |
||
270 |
fun of_show_have qs = if member (op =) qs Show then "show" else "have" |
|
271 |
fun of_thus_hence qs = if member (op =) qs Show then "thus" else "hence" |
|
272 |
||
273 |
fun of_have qs nr = |
|
274 |
if nr > 1 orelse (nr = 1 andalso member (op =) qs Then) then "ultimately " ^ of_show_have qs |
|
275 |
else if nr = 1 orelse member (op =) qs Then then of_thus_hence qs |
|
276 |
else of_show_have qs |
|
277 |
||
278 |
fun add_term term (s, ctxt) = |
|
279 |
(s ^ (term |
|
280 |
|> singleton (Syntax.uncheck_terms ctxt) |
|
| 55213 | 281 |
|> annotate_types_in_term ctxt |
| 55211 | 282 |
|> with_vanilla_print_mode (Syntax.unparse_term ctxt #> Pretty.string_of) |
283 |
|> simplify_spaces |
|
284 |
|> maybe_quote), |
|
285 |
ctxt |> Variable.auto_fixes term) |
|
286 |
||
287 |
fun with_facts none _ [] [] = none |
|
288 |
| with_facts _ some ls ss = some (space_implode " " (map string_of_label ls @ ss)) |
|
289 |
||
290 |
val using_facts = with_facts "" (enclose "using " " ") |
|
|
55214
48a347b40629
better tracing + syntactically correct 'metis' calls
blanchet
parents:
55213
diff
changeset
|
291 |
fun by_facts meth ls ss = "by " ^ with_facts meth (enclose ("(" ^ meth ^ " ") ")") ls ss
|
| 55211 | 292 |
|
293 |
(* Local facts are always passed via "using", which affects "meson" and "metis". This is |
|
294 |
arguably stylistically superior, because it emphasises the structure of the proof. It is also |
|
295 |
more robust w.r.t. preplay: Preplay is performed before chaining of local facts with "hence" |
|
296 |
and "thus" is introduced. See also "tac_of_method" in "Sledgehammer_Isar_Preplay". *) |
|
297 |
fun of_method ls ss Metis_Method = |
|
| 55212 | 298 |
using_facts ls [] ^ by_facts (metis_call type_enc lam_trans) [] ss |
299 |
| of_method ls ss Meson_Method = using_facts ls [] ^ by_facts "meson" [] ss |
|
| 55222 | 300 |
| of_method ls ss meth = using_facts ls ss ^ "by " ^ string_of_proof_method meth |
| 55211 | 301 |
|
302 |
fun of_free (s, T) = |
|
303 |
maybe_quote s ^ " :: " ^ |
|
304 |
maybe_quote (simplify_spaces (with_vanilla_print_mode (Syntax.string_of_typ ctxt) T)) |
|
305 |
||
306 |
fun add_frees xs (s, ctxt) = |
|
307 |
(s ^ space_implode " and " (map of_free xs), ctxt |> register_fixes xs) |
|
308 |
||
309 |
fun add_fix _ [] = I |
|
310 |
| add_fix ind xs = |
|
| 55216 | 311 |
add_str (of_indent ind ^ "fix ") |
| 55211 | 312 |
#> add_frees xs |
| 55216 | 313 |
#> add_str "\n" |
| 55211 | 314 |
|
315 |
fun add_assm ind (l, t) = |
|
| 55216 | 316 |
add_str (of_indent ind ^ "assume " ^ of_label l) |
| 55211 | 317 |
#> add_term t |
| 55216 | 318 |
#> add_str "\n" |
| 55211 | 319 |
|
320 |
fun add_assms ind assms = fold (add_assm ind) assms |
|
321 |
||
322 |
fun of_subproof ind ctxt proof = |
|
323 |
let |
|
324 |
val ind = ind + 1 |
|
325 |
val s = of_proof ind ctxt proof |
|
326 |
val prefix = "{ "
|
|
327 |
val suffix = " }" |
|
328 |
in |
|
329 |
replicate_string (ind * indent_size - size prefix) " " ^ prefix ^ |
|
330 |
String.extract (s, ind * indent_size, SOME (size s - ind * indent_size - 1)) ^ |
|
331 |
suffix ^ "\n" |
|
332 |
end |
|
333 |
and of_subproofs _ _ _ [] = "" |
|
334 |
| of_subproofs ind ctxt qs subproofs = |
|
335 |
(if member (op =) qs Then then of_moreover ind else "") ^ |
|
336 |
space_implode (of_moreover ind) (map (of_subproof ind ctxt) subproofs) |
|
337 |
and add_step_pre ind qs subproofs (s, ctxt) = |
|
338 |
(s ^ of_subproofs ind ctxt qs subproofs ^ of_indent ind, ctxt) |
|
339 |
and add_step ind (Let (t1, t2)) = |
|
| 55216 | 340 |
add_str (of_indent ind ^ "let ") |
341 |
#> add_term t1 #> add_str " = " #> add_term t2 |
|
342 |
#> add_str "\n" |
|
| 55222 | 343 |
| add_step ind (Prove (qs, xs, l, t, subproofs, ((ls, ss), methss as (meth :: _) :: _))) = |
| 55211 | 344 |
add_step_pre ind qs subproofs |
345 |
#> (case xs of |
|
| 55216 | 346 |
[] => add_str (of_have qs (length subproofs) ^ " ") |
347 |
| _ => |
|
348 |
add_str (of_obtain qs (length subproofs) ^ " ") #> add_frees xs #> add_str " where ") |
|
| 55217 | 349 |
#> add_str (of_label l) |
350 |
#> add_term t |
|
351 |
#> add_str (" " ^
|
|
| 55218 | 352 |
of_method (sort_distinct label_ord ls) (sort_distinct string_ord ss) meth ^ |
| 55222 | 353 |
(case comment_of l methss of |
354 |
"" => "" |
|
355 |
| comment => " (* " ^ comment ^ " *)") ^ "\n") |
|
| 55211 | 356 |
and add_steps ind = fold (add_step ind) |
357 |
and of_proof ind ctxt (Proof (xs, assms, steps)) = |
|
358 |
("", ctxt) |> add_fix ind xs |> add_assms ind assms |> add_steps ind steps |> fst
|
|
359 |
in |
|
360 |
(* One-step Metis proofs are pointless; better use the one-liner directly. *) |
|
361 |
(case proof of |
|
362 |
Proof ([], [], []) => "" (* degenerate case: the conjecture is "True" with Z3 *) |
|
363 |
| Proof ([], [], [Prove (_, [], _, _, [], (_, (Metis_Method :: _) :: _))]) => "" |
|
364 |
| _ => |
|
365 |
(if i <> 1 then "prefer " ^ string_of_int i ^ "\n" else "") ^ |
|
366 |
of_indent 0 ^ "proof -\n" ^ of_proof 1 ctxt proof ^ |
|
367 |
of_indent 0 ^ (if n <> 1 then "next" else "qed")) |
|
368 |
end |
|
369 |
||
| 54504 | 370 |
end; |