| author | blanchet | 
| Mon, 07 Mar 2016 23:20:11 +0100 | |
| changeset 62535 | cb262f03ac12 | 
| parent 61756 | 21c2b1f691d1 | 
| child 66020 | a31760eee09d | 
| permissions | -rw-r--r-- | 
| 55202 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 blanchet parents: 
55194diff
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: 
55194diff
changeset | 9 | signature SLEDGEHAMMER_ISAR_PROOF = | 
| 50259 | 10 | sig | 
| 55287 | 11 | type proof_method = Sledgehammer_Proof_Methods.proof_method | 
| 55285 | 12 | |
| 51239 | 13 | type label = string * int | 
| 54816 
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
 blanchet parents: 
54767diff
changeset | 14 | type facts = label list * string list (* local and global facts *) | 
| 50268 | 15 | |
| 51178 | 16 | datatype isar_qualifier = Show | Then | 
| 50268 | 17 | |
| 52454 | 18 | datatype isar_proof = | 
| 54700 | 19 | 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: 
51178diff
changeset | 20 | and isar_step = | 
| 50268 | 21 | Let of term * term | | 
| 55280 | 22 | Prove of isar_qualifier list * (string * typ) list * label * term * isar_proof list | 
| 55299 | 23 | * facts * proof_method list * string | 
| 52592 
8a25b17e3d79
optimize isar-proofs by trying different proof methods
 smolkas parents: 
52590diff
changeset | 24 | |
| 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: 
51178diff
changeset | 25 | val no_label : label | 
| 50268 | 26 | |
| 52592 
8a25b17e3d79
optimize isar-proofs by trying different proof methods
 smolkas parents: 
52590diff
changeset | 27 | val label_ord : label * label -> order | 
| 51998 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 blanchet parents: 
51239diff
changeset | 28 | val string_of_label : label -> string | 
| 57776 
1111a9a328fe
rationalized sorting of facts -- so that preplaying (almost always) coincides with the real thing, preventing odd failures
 blanchet parents: 
57765diff
changeset | 29 | val sort_facts : facts -> facts | 
| 52592 
8a25b17e3d79
optimize isar-proofs by trying different proof methods
 smolkas parents: 
52590diff
changeset | 30 | |
| 55260 | 31 | val steps_of_isar_proof : isar_proof -> isar_step list | 
| 55223 
3c593bad6b31
generalized preplaying infrastructure to store various results for various methods
 blanchet parents: 
55222diff
changeset | 32 | val label_of_isar_step : isar_step -> label option | 
| 55279 | 33 | val facts_of_isar_step : isar_step -> facts | 
| 34 | val proof_methods_of_isar_step : isar_step -> proof_method 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: 
51178diff
changeset | 35 | |
| 54765 | 36 | val fold_isar_steps : (isar_step -> 'a -> 'a) -> isar_step list -> 'a -> 'a | 
| 55212 | 37 | val map_isar_steps : (isar_step -> isar_step) -> isar_proof -> isar_proof | 
| 38 | val add_isar_steps : isar_step list -> int -> int | |
| 52454 | 39 | |
| 55212 | 40 | structure Canonical_Label_Tab : TABLE | 
| 52556 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 41 | |
| 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 42 | val canonical_label_ord : (label * label) -> order | 
| 55220 | 43 | |
| 55299 | 44 | val comment_isar_proof : (label -> proof_method list -> string) -> isar_proof -> isar_proof | 
| 55220 | 45 | val chain_isar_proof : isar_proof -> isar_proof | 
| 46 | val kill_useless_labels_in_isar_proof : isar_proof -> isar_proof | |
| 55213 | 47 | val relabel_isar_proof_canonically : isar_proof -> isar_proof | 
| 55282 | 48 | val relabel_isar_proof_nicely : isar_proof -> isar_proof | 
| 57792 | 49 | val rationalize_obtains_in_isar_proofs : Proof.context -> isar_proof -> isar_proof | 
| 52556 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 50 | |
| 55299 | 51 | val string_of_isar_proof : Proof.context -> int -> int -> isar_proof -> string | 
| 54504 | 52 | end; | 
| 50259 | 53 | |
| 55202 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 blanchet parents: 
55194diff
changeset | 54 | structure Sledgehammer_Isar_Proof : SLEDGEHAMMER_ISAR_PROOF = | 
| 50259 | 55 | struct | 
| 56 | ||
| 55211 | 57 | open ATP_Util | 
| 58 | open ATP_Proof | |
| 59 | open ATP_Problem_Generate | |
| 60 | open ATP_Proof_Reconstruct | |
| 61 | open Sledgehammer_Util | |
| 55287 | 62 | open Sledgehammer_Proof_Methods | 
| 55211 | 63 | open Sledgehammer_Isar_Annotate | 
| 64 | ||
| 50259 | 65 | type label = string * int | 
| 54534 | 66 | type facts = label list * string list (* local and global facts *) | 
| 50259 | 67 | |
| 51178 | 68 | datatype isar_qualifier = Show | Then | 
| 50259 | 69 | |
| 52454 | 70 | datatype isar_proof = | 
| 54700 | 71 | 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: 
51178diff
changeset | 72 | and isar_step = | 
| 50259 | 73 | Let of term * term | | 
| 55280 | 74 | Prove of isar_qualifier list * (string * typ) list * label * term * isar_proof list | 
| 55299 | 75 | * facts * proof_method list * string | 
| 52592 
8a25b17e3d79
optimize isar-proofs by trying different proof methods
 smolkas parents: 
52590diff
changeset | 76 | |
| 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: 
51178diff
changeset | 77 | val no_label = ("", ~1)
 | 
| 50259 | 78 | |
| 57776 
1111a9a328fe
rationalized sorting of facts -- so that preplaying (almost always) coincides with the real thing, preventing odd failures
 blanchet parents: 
57765diff
changeset | 79 | (* cf. "label_ord" below *) | 
| 
1111a9a328fe
rationalized sorting of facts -- so that preplaying (almost always) coincides with the real thing, preventing odd failures
 blanchet parents: 
57765diff
changeset | 80 | val assume_prefix = "a" | 
| 
1111a9a328fe
rationalized sorting of facts -- so that preplaying (almost always) coincides with the real thing, preventing odd failures
 blanchet parents: 
57765diff
changeset | 81 | val have_prefix = "f" | 
| 
1111a9a328fe
rationalized sorting of facts -- so that preplaying (almost always) coincides with the real thing, preventing odd failures
 blanchet parents: 
57765diff
changeset | 82 | |
| 
1111a9a328fe
rationalized sorting of facts -- so that preplaying (almost always) coincides with the real thing, preventing odd failures
 blanchet parents: 
57765diff
changeset | 83 | fun label_ord ((s1, i1), (s2, i2)) = | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
58928diff
changeset | 84 | (case int_ord (apply2 String.size (s1, s2)) of | 
| 57776 
1111a9a328fe
rationalized sorting of facts -- so that preplaying (almost always) coincides with the real thing, preventing odd failures
 blanchet parents: 
57765diff
changeset | 85 | EQUAL => | 
| 
1111a9a328fe
rationalized sorting of facts -- so that preplaying (almost always) coincides with the real thing, preventing odd failures
 blanchet parents: 
57765diff
changeset | 86 | (case string_ord (s1, s2) of | 
| 
1111a9a328fe
rationalized sorting of facts -- so that preplaying (almost always) coincides with the real thing, preventing odd failures
 blanchet parents: 
57765diff
changeset | 87 | EQUAL => int_ord (i1, i2) | 
| 
1111a9a328fe
rationalized sorting of facts -- so that preplaying (almost always) coincides with the real thing, preventing odd failures
 blanchet parents: 
57765diff
changeset | 88 | | ord => ord (* "assume" before "have" *)) | 
| 
1111a9a328fe
rationalized sorting of facts -- so that preplaying (almost always) coincides with the real thing, preventing odd failures
 blanchet parents: 
57765diff
changeset | 89 | | ord => ord) | 
| 52556 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 90 | |
| 51998 
f732a674db1b
renamed Sledgehammer functions with 'for' in their names to 'of'
 blanchet parents: 
51239diff
changeset | 91 | fun string_of_label (s, num) = s ^ string_of_int num | 
| 50259 | 92 | |
| 57776 
1111a9a328fe
rationalized sorting of facts -- so that preplaying (almost always) coincides with the real thing, preventing odd failures
 blanchet parents: 
57765diff
changeset | 93 | (* Put the nearest local label first, since it's the most likely to be replaced by a "hence". | 
| 
1111a9a328fe
rationalized sorting of facts -- so that preplaying (almost always) coincides with the real thing, preventing odd failures
 blanchet parents: 
57765diff
changeset | 94 | (Some preplaying proof methods, e.g. "blast", react poorly to fact reorderings.) *) | 
| 
1111a9a328fe
rationalized sorting of facts -- so that preplaying (almost always) coincides with the real thing, preventing odd failures
 blanchet parents: 
57765diff
changeset | 95 | fun sort_facts (lfs, gfs) = (sort (label_ord o swap) lfs, sort string_ord gfs) | 
| 
1111a9a328fe
rationalized sorting of facts -- so that preplaying (almost always) coincides with the real thing, preventing odd failures
 blanchet parents: 
57765diff
changeset | 96 | |
| 55260 | 97 | fun steps_of_isar_proof (Proof (_, _, steps)) = steps | 
| 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: 
51178diff
changeset | 98 | |
| 55299 | 99 | fun label_of_isar_step (Prove (_, _, l, _, _, _, _, _)) = SOME l | 
| 55223 
3c593bad6b31
generalized preplaying infrastructure to store various results for various methods
 blanchet parents: 
55222diff
changeset | 100 | | label_of_isar_step _ = NONE | 
| 51178 | 101 | |
| 55299 | 102 | fun subproofs_of_isar_step (Prove (_, _, _, _, subs, _, _, _)) = subs | 
| 55281 | 103 | | subproofs_of_isar_step _ = [] | 
| 52454 | 104 | |
| 55299 | 105 | fun facts_of_isar_step (Prove (_, _, _, _, _, facts, _, _)) = facts | 
| 55279 | 106 | | facts_of_isar_step _ = ([], []) | 
| 107 | ||
| 55299 | 108 | fun proof_methods_of_isar_step (Prove (_, _, _, _, _, _, meths, _)) = meths | 
| 55279 | 109 | | proof_methods_of_isar_step _ = [] | 
| 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: 
51178diff
changeset | 110 | |
| 55212 | 111 | fun fold_isar_step f step = | 
| 55281 | 112 | fold (steps_of_isar_proof #> fold_isar_steps f) (subproofs_of_isar_step step) #> f step | 
| 55212 | 113 | and fold_isar_steps f = fold (fold_isar_step f) | 
| 52454 | 114 | |
| 54765 | 115 | fun map_isar_steps f = | 
| 52592 
8a25b17e3d79
optimize isar-proofs by trying different proof methods
 smolkas parents: 
52590diff
changeset | 116 | let | 
| 55212 | 117 | fun map_proof (Proof (fix, assms, steps)) = Proof (fix, assms, map map_step steps) | 
| 118 | and map_step (step as Let _) = f step | |
| 55299 | 119 | | map_step (Prove (qs, xs, l, t, subs, facts, meths, comment)) = | 
| 120 | f (Prove (qs, xs, l, t, map map_proof subs, facts, meths, comment)) | |
| 55212 | 121 | in map_proof end | 
| 52592 
8a25b17e3d79
optimize isar-proofs by trying different proof methods
 smolkas parents: 
52590diff
changeset | 122 | |
| 55212 | 123 | 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: 
52454diff
changeset | 124 | |
| 55211 | 125 | (* canonical proof labels: 1, 2, 3, ... in post traversal order *) | 
| 52557 | 126 | 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: 
52454diff
changeset | 127 | |
| 55212 | 128 | structure Canonical_Label_Tab = Table( | 
| 52556 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 129 | type key = label | 
| 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 130 | val ord = canonical_label_ord) | 
| 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 131 | |
| 55299 | 132 | fun comment_isar_step comment_of (Prove (qs, xs, l, t, subs, facts, meths, _)) = | 
| 133 | Prove (qs, xs, l, t, subs, facts, meths, comment_of l meths) | |
| 134 | | comment_isar_step _ step = step | |
| 135 | fun comment_isar_proof comment_of = map_isar_steps (comment_isar_step comment_of) | |
| 136 | ||
| 55220 | 137 | fun chain_qs_lfs NONE lfs = ([], lfs) | 
| 138 | | chain_qs_lfs (SOME l0) lfs = | |
| 57655 | 139 | if member (op =) lfs l0 then ([Then], remove (op =) l0 lfs) else ([], lfs) | 
| 140 | ||
| 57286 
4868ec62f533
don't generate Isar proof skeleton for what amounts to a one-line proof
 blanchet parents: 
57154diff
changeset | 141 | fun chain_isar_step lbl (Prove (qs, xs, l, t, subs, (lfs, gfs), meths, comment)) = | 
| 55220 | 142 | let val (qs', lfs) = chain_qs_lfs lbl lfs in | 
| 55299 | 143 | Prove (qs' @ qs, xs, l, t, map chain_isar_proof subs, (lfs, gfs), meths, comment) | 
| 55220 | 144 | end | 
| 145 | | chain_isar_step _ step = step | |
| 146 | and chain_isar_steps _ [] = [] | |
| 57655 | 147 | | chain_isar_steps prev (i :: is) = | 
| 55223 
3c593bad6b31
generalized preplaying infrastructure to store various results for various methods
 blanchet parents: 
55222diff
changeset | 148 | chain_isar_step prev i :: chain_isar_steps (label_of_isar_step i) is | 
| 55220 | 149 | and chain_isar_proof (Proof (fix, assms, steps)) = | 
| 150 | Proof (fix, assms, chain_isar_steps (try (List.last #> fst) assms) steps) | |
| 151 | ||
| 152 | fun kill_useless_labels_in_isar_proof proof = | |
| 153 | let | |
| 154 | val used_ls = | |
| 55279 | 155 | fold_isar_steps (facts_of_isar_step #> fst #> union (op =)) (steps_of_isar_proof proof) [] | 
| 55220 | 156 | |
| 157 | fun kill_label l = if member (op =) used_ls l then l else no_label | |
| 158 | ||
| 55299 | 159 | fun kill_step (Prove (qs, xs, l, t, subs, facts, meths, comment)) = | 
| 160 | Prove (qs, xs, kill_label l, t, map kill_proof subs, facts, meths, comment) | |
| 55220 | 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: 
52454diff
changeset | 169 | let | 
| 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
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: 
52454diff
changeset | 172 | |
| 55299 | 173 | fun relabel_step (Prove (qs, fix, l, t, subs, (lfs, gfs), meths, comment)) | 
| 174 | (accum as (_, subst)) = | |
| 54534 | 175 | let | 
| 55281 | 176 | val lfs' = maps (the_list o AList.lookup (op =) subst) lfs | 
| 177 | val ((subs', l'), accum') = accum | |
| 178 | |> fold_map relabel_proof subs | |
| 179 | ||>> next_label l | |
| 54534 | 180 | in | 
| 55299 | 181 | (Prove (qs, fix, l', t, subs', (lfs', gfs), meths, comment), accum') | 
| 54534 | 182 | end | 
| 55281 | 183 | | relabel_step step accum = (step, accum) | 
| 184 | and relabel_proof (Proof (fix, assms, steps)) = | |
| 185 | fold_map (fn (l, t) => next_label l #> apfst (rpair t)) assms | |
| 186 | ##>> fold_map relabel_step steps | |
| 187 | #>> (fn (assms, steps) => Proof (fix, assms, steps)) | |
| 52556 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 188 | in | 
| 55279 | 189 | fst (relabel_proof proof (0, [])) | 
| 52556 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 190 | end | 
| 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 191 | |
| 55282 | 192 | val relabel_isar_proof_nicely = | 
| 55220 | 193 | let | 
| 55281 | 194 | fun next_label depth prefix l (accum as (next, subst)) = | 
| 55220 | 195 | if l = no_label then | 
| 55281 | 196 | (l, accum) | 
| 55220 | 197 | else | 
| 198 | let val l' = (replicate_string (depth + 1) prefix, next) in | |
| 55281 | 199 | (l', (next + 1, (l, l') :: subst)) | 
| 55220 | 200 | end | 
| 201 | ||
| 55299 | 202 | fun relabel_step depth (Prove (qs, xs, l, t, subs, (lfs, gfs), meths, comment)) | 
| 203 | (accum as (_, subst)) = | |
| 55220 | 204 | let | 
| 55281 | 205 | val lfs' = maps (the_list o AList.lookup (op =) subst) lfs | 
| 55309 | 206 | val (l', accum' as (_, subst')) = next_label depth have_prefix l accum | 
| 55281 | 207 | val subs' = map (relabel_proof subst' (depth + 1)) subs | 
| 55220 | 208 | in | 
| 55299 | 209 | (Prove (qs, xs, l', t, subs', (lfs', gfs), meths, comment), accum') | 
| 55220 | 210 | end | 
| 55281 | 211 | | relabel_step _ step accum = (step, accum) | 
| 55220 | 212 | and relabel_proof subst depth (Proof (fix, assms, steps)) = | 
| 55281 | 213 | (1, subst) | 
| 214 | |> fold_map (fn (l, t) => next_label depth assume_prefix l #> apfst (rpair t)) assms | |
| 215 | ||>> fold_map (relabel_step depth) steps | |
| 216 | |> (fn ((assms, steps), _) => Proof (fix, assms, steps)) | |
| 55220 | 217 | in | 
| 218 | relabel_proof [] 0 | |
| 219 | end | |
| 220 | ||
| 57793 | 221 | fun stutter_single_letter s = String.extract (s, 0, SOME 1) ^ s | 
| 222 | ||
| 57792 | 223 | fun rationalize_obtains_in_isar_proofs ctxt = | 
| 224 | let | |
| 225 | fun rename_obtains xs (subst, ctxt) = | |
| 226 | let | |
| 227 | val Ts = map snd xs | |
| 57793 | 228 | val new_names0 = map (stutter_single_letter o var_name_of_typ o body_type) Ts | 
| 57792 | 229 | val (new_names, ctxt') = Variable.variant_fixes new_names0 ctxt | 
| 230 | val ys = map2 pair new_names Ts | |
| 231 | in | |
| 232 | (ys, ((map Free xs ~~ map Free ys) @ subst, ctxt')) | |
| 233 | end | |
| 234 | ||
| 235 | fun rationalize_step (Prove (qs, xs, l, t, subs, facts, meths, comment)) subst_ctxt = | |
| 236 | let | |
| 237 | val (xs', subst_ctxt' as (subst', _)) = rename_obtains xs subst_ctxt | |
| 238 | val t' = subst_atomic subst' t | |
| 58475 
4508b6bff671
rename skolem symbols in the negative case as well
 blanchet parents: 
58087diff
changeset | 239 | val subs' = map (rationalize_proof false subst_ctxt') subs | 
| 57792 | 240 | in | 
| 241 | (Prove (qs, xs', l, t', subs', facts, meths, comment), subst_ctxt') | |
| 242 | end | |
| 58476 
6ade4c7109a8
make sure no '__' suffixes make it until Isar proof
 blanchet parents: 
58475diff
changeset | 243 | and rationalize_proof outer (subst_ctxt as (subst, ctxt)) (Proof (fix, assms, steps)) = | 
| 58475 
4508b6bff671
rename skolem symbols in the negative case as well
 blanchet parents: 
58087diff
changeset | 244 | let | 
| 
4508b6bff671
rename skolem symbols in the negative case as well
 blanchet parents: 
58087diff
changeset | 245 | val (fix', subst_ctxt' as (subst', _)) = | 
| 58476 
6ade4c7109a8
make sure no '__' suffixes make it until Isar proof
 blanchet parents: 
58475diff
changeset | 246 | if outer then | 
| 
6ade4c7109a8
make sure no '__' suffixes make it until Isar proof
 blanchet parents: 
58475diff
changeset | 247 | (* last call: eliminate any remaining skolem names (from SMT proofs) *) | 
| 
6ade4c7109a8
make sure no '__' suffixes make it until Isar proof
 blanchet parents: 
58475diff
changeset | 248 | (fix, (subst @ map (fn x => (Free (apfst Name.skolem x), Free x)) fix, ctxt)) | 
| 
6ade4c7109a8
make sure no '__' suffixes make it until Isar proof
 blanchet parents: 
58475diff
changeset | 249 | else | 
| 
6ade4c7109a8
make sure no '__' suffixes make it until Isar proof
 blanchet parents: 
58475diff
changeset | 250 | rename_obtains fix subst_ctxt | 
| 58475 
4508b6bff671
rename skolem symbols in the negative case as well
 blanchet parents: 
58087diff
changeset | 251 | in | 
| 
4508b6bff671
rename skolem symbols in the negative case as well
 blanchet parents: 
58087diff
changeset | 252 | Proof (fix', map (apsnd (subst_atomic subst')) assms, | 
| 
4508b6bff671
rename skolem symbols in the negative case as well
 blanchet parents: 
58087diff
changeset | 253 | fst (fold_map rationalize_step steps subst_ctxt')) | 
| 
4508b6bff671
rename skolem symbols in the negative case as well
 blanchet parents: 
58087diff
changeset | 254 | end | 
| 57792 | 255 | in | 
| 58475 
4508b6bff671
rename skolem symbols in the negative case as well
 blanchet parents: 
58087diff
changeset | 256 | rationalize_proof true ([], ctxt) | 
| 57792 | 257 | end | 
| 258 | ||
| 58087 
32d3fa94ebb4
use 'thesis' only if it expands to the right thing (it won't after an 'unfolding', for example)
 blanchet parents: 
58082diff
changeset | 259 | val thesis_var = ((Auto_Bind.thesisN, 0), HOLogic.boolT) | 
| 
32d3fa94ebb4
use 'thesis' only if it expands to the right thing (it won't after an 'unfolding', for example)
 blanchet parents: 
58082diff
changeset | 260 | |
| 
32d3fa94ebb4
use 'thesis' only if it expands to the right thing (it won't after an 'unfolding', for example)
 blanchet parents: 
58082diff
changeset | 261 | fun is_thesis ctxt t = | 
| 
32d3fa94ebb4
use 'thesis' only if it expands to the right thing (it won't after an 'unfolding', for example)
 blanchet parents: 
58082diff
changeset | 262 | (case Vartab.lookup (Variable.binds_of ctxt) (fst thesis_var) of | 
| 
32d3fa94ebb4
use 'thesis' only if it expands to the right thing (it won't after an 'unfolding', for example)
 blanchet parents: 
58082diff
changeset | 263 | SOME (_, t') => HOLogic.mk_Trueprop t' aconv t | 
| 
32d3fa94ebb4
use 'thesis' only if it expands to the right thing (it won't after an 'unfolding', for example)
 blanchet parents: 
58082diff
changeset | 264 | | NONE => false) | 
| 
32d3fa94ebb4
use 'thesis' only if it expands to the right thing (it won't after an 'unfolding', for example)
 blanchet parents: 
58082diff
changeset | 265 | |
| 55211 | 266 | val indent_size = 2 | 
| 267 | ||
| 56985 
82c83978fbd9
correctly add extra facts to lemmas (cf. conjecture and hypotheses) in Z3 Isar proofs
 blanchet parents: 
56983diff
changeset | 268 | fun string_of_isar_proof ctxt0 i n proof = | 
| 55211 | 269 | let | 
| 58928 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58476diff
changeset | 270 | val keywords = Thy_Header.get_keywords' ctxt0 | 
| 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58476diff
changeset | 271 | |
| 55211 | 272 | (* Make sure only type constraints inserted by the type annotation code are printed. *) | 
| 56985 
82c83978fbd9
correctly add extra facts to lemmas (cf. conjecture and hypotheses) in Z3 Isar proofs
 blanchet parents: 
56983diff
changeset | 273 | val ctxt = ctxt0 | 
| 
82c83978fbd9
correctly add extra facts to lemmas (cf. conjecture and hypotheses) in Z3 Isar proofs
 blanchet parents: 
56983diff
changeset | 274 | |> Config.put show_markup false | 
| 
82c83978fbd9
correctly add extra facts to lemmas (cf. conjecture and hypotheses) in Z3 Isar proofs
 blanchet parents: 
56983diff
changeset | 275 | |> Config.put Printer.show_type_emphasis false | 
| 
82c83978fbd9
correctly add extra facts to lemmas (cf. conjecture and hypotheses) in Z3 Isar proofs
 blanchet parents: 
56983diff
changeset | 276 | |> Config.put show_types false | 
| 
82c83978fbd9
correctly add extra facts to lemmas (cf. conjecture and hypotheses) in Z3 Isar proofs
 blanchet parents: 
56983diff
changeset | 277 | |> Config.put show_sorts false | 
| 
82c83978fbd9
correctly add extra facts to lemmas (cf. conjecture and hypotheses) in Z3 Isar proofs
 blanchet parents: 
56983diff
changeset | 278 | |> Config.put show_consts false | 
| 55211 | 279 | |
| 55216 | 280 | fun add_str s' = apfst (suffix s') | 
| 55211 | 281 | |
| 282 | fun of_indent ind = replicate_string (ind * indent_size) " " | |
| 283 | fun of_moreover ind = of_indent ind ^ "moreover\n" | |
| 284 | fun of_label l = if l = no_label then "" else string_of_label l ^ ": " | |
| 285 | ||
| 286 | fun of_obtain qs nr = | |
| 287 | (if nr > 1 orelse (nr = 1 andalso member (op =) qs Then) then "ultimately " | |
| 288 | else if nr = 1 orelse member (op =) qs Then then "then " | |
| 289 | else "") ^ "obtain" | |
| 290 | ||
| 291 | fun of_show_have qs = if member (op =) qs Show then "show" else "have" | |
| 61756 | 292 | fun of_thus_hence qs = if member (op =) qs Show then "then show" else "then have" | 
| 55211 | 293 | |
| 294 | fun of_have qs nr = | |
| 295 | if nr > 1 orelse (nr = 1 andalso member (op =) qs Then) then "ultimately " ^ of_show_have qs | |
| 296 | else if nr = 1 orelse member (op =) qs Then then of_thus_hence qs | |
| 297 | else of_show_have qs | |
| 298 | ||
| 299 | fun add_term term (s, ctxt) = | |
| 300 | (s ^ (term | |
| 301 | |> singleton (Syntax.uncheck_terms ctxt) | |
| 55213 | 302 | |> annotate_types_in_term ctxt | 
| 55211 | 303 | |> with_vanilla_print_mode (Syntax.unparse_term ctxt #> Pretty.string_of) | 
| 304 | |> simplify_spaces | |
| 58928 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58476diff
changeset | 305 | |> maybe_quote keywords), | 
| 57669 | 306 | ctxt |> perhaps (try (Variable.auto_fixes term))) | 
| 55211 | 307 | |
| 56983 
132142089ea6
use 'simp add:' syntax in Sledgehammer rather than 'using'
 blanchet parents: 
56081diff
changeset | 308 | fun using_facts [] [] = "" | 
| 
132142089ea6
use 'simp add:' syntax in Sledgehammer rather than 'using'
 blanchet parents: 
56081diff
changeset | 309 | | using_facts ls ss = enclose "using " " " (space_implode " " (map string_of_label ls @ ss)) | 
| 55257 | 310 | |
| 55211 | 311 | (* Local facts are always passed via "using", which affects "meson" and "metis". This is | 
| 312 | arguably stylistically superior, because it emphasises the structure of the proof. It is also | |
| 61756 | 313 | more robust w.r.t. preplay: Preplay is performed before chaining of local facts with "then" | 
| 314 | is introduced. See also "tac_of_method" in "Sledgehammer_Isar_Preplay". *) | |
| 57287 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 315 | fun of_method ls ss meth = | 
| 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 316 | let val direct = is_proof_method_direct meth in | 
| 55281 | 317 | using_facts ls (if direct then [] else ss) ^ | 
| 57287 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 318 | "by " ^ string_of_proof_method ctxt (if direct then ss else []) meth | 
| 55281 | 319 | end | 
| 55211 | 320 | |
| 321 | fun of_free (s, T) = | |
| 58928 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58476diff
changeset | 322 | maybe_quote keywords s ^ " :: " ^ | 
| 
23d0ffd48006
plain value Keywords.keywords, which might be used outside theory for bootstrap purposes;
 wenzelm parents: 
58476diff
changeset | 323 | maybe_quote keywords (simplify_spaces (with_vanilla_print_mode (Syntax.string_of_typ ctxt) T)) | 
| 55211 | 324 | |
| 325 | fun add_frees xs (s, ctxt) = | |
| 57669 | 326 | (s ^ space_implode " and " (map of_free xs), ctxt |> fold Variable.auto_fixes (map Free xs)) | 
| 55211 | 327 | |
| 328 | fun add_fix _ [] = I | |
| 55281 | 329 | | add_fix ind xs = add_str (of_indent ind ^ "fix ") #> add_frees xs #> add_str "\n" | 
| 55211 | 330 | |
| 331 | fun add_assm ind (l, t) = | |
| 55281 | 332 | add_str (of_indent ind ^ "assume " ^ of_label l) #> add_term t #> add_str "\n" | 
| 55211 | 333 | |
| 334 | fun of_subproof ind ctxt proof = | |
| 335 | let | |
| 336 | val ind = ind + 1 | |
| 337 | val s = of_proof ind ctxt proof | |
| 338 |         val prefix = "{ "
 | |
| 339 | val suffix = " }" | |
| 340 | in | |
| 341 | replicate_string (ind * indent_size - size prefix) " " ^ prefix ^ | |
| 342 | String.extract (s, ind * indent_size, SOME (size s - ind * indent_size - 1)) ^ | |
| 343 | suffix ^ "\n" | |
| 344 | end | |
| 345 | and of_subproofs _ _ _ [] = "" | |
| 55281 | 346 | | of_subproofs ind ctxt qs subs = | 
| 55211 | 347 | (if member (op =) qs Then then of_moreover ind else "") ^ | 
| 55281 | 348 | space_implode (of_moreover ind) (map (of_subproof ind ctxt) subs) | 
| 349 | and add_step_pre ind qs subs (s, ctxt) = | |
| 350 | (s ^ of_subproofs ind ctxt qs subs ^ of_indent ind, ctxt) | |
| 55211 | 351 | and add_step ind (Let (t1, t2)) = | 
| 57776 
1111a9a328fe
rationalized sorting of facts -- so that preplaying (almost always) coincides with the real thing, preventing odd failures
 blanchet parents: 
57765diff
changeset | 352 | add_str (of_indent ind ^ "let ") #> add_term t1 #> add_str " = " #> add_term t2 | 
| 
1111a9a328fe
rationalized sorting of facts -- so that preplaying (almost always) coincides with the real thing, preventing odd failures
 blanchet parents: 
57765diff
changeset | 353 | #> add_str "\n" | 
| 55309 | 354 | | add_step ind (Prove (qs, xs, l, t, subs, (ls, ss), meth :: _, comment)) = | 
| 55281 | 355 | add_step_pre ind qs subs | 
| 55211 | 356 | #> (case xs of | 
| 57776 
1111a9a328fe
rationalized sorting of facts -- so that preplaying (almost always) coincides with the real thing, preventing odd failures
 blanchet parents: 
57765diff
changeset | 357 | [] => add_str (of_have qs (length subs) ^ " ") | 
| 58071 
62ec5b1d2f2f
three-line 'obtain' format for generated Isar proofs
 blanchet parents: 
58054diff
changeset | 358 | | _ => | 
| 
62ec5b1d2f2f
three-line 'obtain' format for generated Isar proofs
 blanchet parents: 
58054diff
changeset | 359 | add_str (of_obtain qs (length subs) ^ " ") | 
| 
62ec5b1d2f2f
three-line 'obtain' format for generated Isar proofs
 blanchet parents: 
58054diff
changeset | 360 | #> add_frees xs | 
| 
62ec5b1d2f2f
three-line 'obtain' format for generated Isar proofs
 blanchet parents: 
58054diff
changeset | 361 |             #> add_str (" where\n" ^ of_indent (ind + 1)))
 | 
| 55217 | 362 | #> add_str (of_label l) | 
| 58087 
32d3fa94ebb4
use 'thesis' only if it expands to the right thing (it won't after an 'unfolding', for example)
 blanchet parents: 
58082diff
changeset | 363 | #> add_term (if is_thesis ctxt t then HOLogic.mk_Trueprop (Var thesis_var) else t) | 
| 58054 
1d9edd486479
reintroduced two-line-per-inference Isar proof format
 blanchet parents: 
57793diff
changeset | 364 |         #> add_str ("\n" ^ of_indent (ind + 1) ^ of_method ls ss meth ^
 | 
| 57776 
1111a9a328fe
rationalized sorting of facts -- so that preplaying (almost always) coincides with the real thing, preventing odd failures
 blanchet parents: 
57765diff
changeset | 365 | (if comment = "" then "" else " (* " ^ comment ^ " *)") ^ "\n") | 
| 55211 | 366 | and add_steps ind = fold (add_step ind) | 
| 367 | and of_proof ind ctxt (Proof (xs, assms, steps)) = | |
| 55281 | 368 |       ("", ctxt)
 | 
| 369 | |> add_fix ind xs | |
| 370 | |> fold (add_assm ind) assms | |
| 371 | |> add_steps ind steps | |
| 372 | |> fst | |
| 55211 | 373 | in | 
| 57286 
4868ec62f533
don't generate Isar proof skeleton for what amounts to a one-line proof
 blanchet parents: 
57154diff
changeset | 374 | (if i <> 1 then "prefer " ^ string_of_int i ^ "\n" else "") ^ | 
| 57287 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 375 | of_indent 0 ^ "proof -\n" ^ of_proof 1 ctxt proof ^ | 
| 
68aa585269ac
given two one-liners, only show the best of the two
 blanchet parents: 
57286diff
changeset | 376 | of_indent 0 ^ (if n = 1 then "qed" else "next") | 
| 55211 | 377 | end | 
| 378 | ||
| 54504 | 379 | end; |