| author | wenzelm | 
| Wed, 05 Mar 2014 16:13:24 +0100 | |
| changeset 55919 | 2eb8c13339a5 | 
| parent 55452 | 29ec8680e61f | 
| child 57054 | fed0329ea8e2 | 
| permissions | -rw-r--r-- | 
| 55202 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 blanchet parents: 
55183diff
changeset | 1 | (* Title: HOL/Tools/Sledgehammer/sledgehammer_isar_compress.ML | 
| 54813 | 2 | Author: Steffen Juilf Smolka, TU Muenchen | 
| 50263 | 3 | Author: Jasmin Blanchette, TU Muenchen | 
| 4 | ||
| 55202 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 blanchet parents: 
55183diff
changeset | 5 | Compression of Isar proofs by merging steps. | 
| 54712 | 6 | Only proof steps using the same proof method are merged. | 
| 50263 | 7 | *) | 
| 8 | ||
| 55202 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 blanchet parents: 
55183diff
changeset | 9 | signature SLEDGEHAMMER_ISAR_COMPRESS = | 
| 50259 | 10 | sig | 
| 55202 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 blanchet parents: 
55183diff
changeset | 11 | type isar_proof = Sledgehammer_Isar_Proof.isar_proof | 
| 55213 | 12 | type isar_preplay_data = Sledgehammer_Isar_Preplay.isar_preplay_data | 
| 52556 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 13 | |
| 55452 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55333diff
changeset | 14 | val compress_isar_proof : Proof.context -> bool -> real -> Time.time -> | 
| 55329 
3c2dbd2e221f
more generous Isar proof compression -- try to remove failing steps
 blanchet parents: 
55328diff
changeset | 15 | isar_preplay_data Unsynchronized.ref -> isar_proof -> isar_proof | 
| 54504 | 16 | end; | 
| 52556 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 17 | |
| 55202 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 blanchet parents: 
55183diff
changeset | 18 | structure Sledgehammer_Isar_Compress : SLEDGEHAMMER_ISAR_COMPRESS = | 
| 50259 | 19 | struct | 
| 20 | ||
| 50265 | 21 | open Sledgehammer_Util | 
| 55287 | 22 | open Sledgehammer_Proof_Methods | 
| 55202 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 blanchet parents: 
55183diff
changeset | 23 | open Sledgehammer_Isar_Proof | 
| 
824c48a539c9
renamed many Sledgehammer ML files to clarify structure
 blanchet parents: 
55183diff
changeset | 24 | open Sledgehammer_Isar_Preplay | 
| 50259 | 25 | |
| 52556 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 26 | fun collect_successors steps lbls = | 
| 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 27 | let | 
| 55329 
3c2dbd2e221f
more generous Isar proof compression -- try to remove failing steps
 blanchet parents: 
55328diff
changeset | 28 | fun collect_steps _ (accum as ([], _)) = accum | 
| 55263 | 29 | | collect_steps [] accum = accum | 
| 30 | | collect_steps (step :: steps) accum = collect_steps steps (collect_step step accum) | |
| 55329 
3c2dbd2e221f
more generous Isar proof compression -- try to remove failing steps
 blanchet parents: 
55328diff
changeset | 31 | and collect_step (step as Prove (_, _, l, _, subproofs, _, _, _)) x = | 
| 55263 | 32 | (case collect_subproofs subproofs x of | 
| 55329 
3c2dbd2e221f
more generous Isar proof compression -- try to remove failing steps
 blanchet parents: 
55328diff
changeset | 33 | (accum as ([], _)) => accum | 
| 54712 | 34 | | accum as (l' :: lbls', accu) => if l = l' then (lbls', step :: accu) else accum) | 
| 55329 
3c2dbd2e221f
more generous Isar proof compression -- try to remove failing steps
 blanchet parents: 
55328diff
changeset | 35 | | collect_step _ accum = accum | 
| 
3c2dbd2e221f
more generous Isar proof compression -- try to remove failing steps
 blanchet parents: 
55328diff
changeset | 36 | and collect_subproofs [] accum = accum | 
| 
3c2dbd2e221f
more generous Isar proof compression -- try to remove failing steps
 blanchet parents: 
55328diff
changeset | 37 | | collect_subproofs (proof :: subproofs) accum = | 
| 
3c2dbd2e221f
more generous Isar proof compression -- try to remove failing steps
 blanchet parents: 
55328diff
changeset | 38 | (case collect_steps (steps_of_isar_proof proof) accum of | 
| 54712 | 39 | accum as ([], _) => accum | 
| 55263 | 40 | | accum => collect_subproofs subproofs accum) | 
| 52556 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 41 | in | 
| 55313 | 42 | rev (snd (collect_steps steps (lbls, []))) | 
| 52556 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 43 | end | 
| 50259 | 44 | |
| 55332 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 45 | fun update_steps updates steps = | 
| 52556 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 46 | let | 
| 55263 | 47 | fun update_steps [] updates = ([], updates) | 
| 48 | | update_steps steps [] = (steps, []) | |
| 49 | | update_steps (step :: steps) updates = update_step step (update_steps steps updates) | |
| 50 | and update_step step (steps, []) = (step :: steps, []) | |
| 55299 | 51 | | update_step (Prove (qs, xs, l, t, subproofs, facts, meths, comment)) | 
| 52 | (steps, | |
| 53 | updates as Prove (qs', xs', l', t', subproofs', facts', meths', comment') :: updates') = | |
| 54 | (if l = l' then | |
| 55 | update_subproofs subproofs' updates' | |
| 56 | |>> (fn subproofs' => Prove (qs', xs', l', t', subproofs', facts', meths', comment')) | |
| 57 | else | |
| 58 | update_subproofs subproofs updates | |
| 59 | |>> (fn subproofs => Prove (qs, xs, l, t, subproofs, facts, meths, comment))) | |
| 60 | |>> (fn step => step :: steps) | |
| 55329 
3c2dbd2e221f
more generous Isar proof compression -- try to remove failing steps
 blanchet parents: 
55328diff
changeset | 61 | | update_step step (steps, updates) = (step :: steps, updates) | 
| 55263 | 62 | and update_subproofs [] updates = ([], updates) | 
| 63 | | update_subproofs steps [] = (steps, []) | |
| 64 | | update_subproofs (proof :: subproofs) updates = | |
| 65 | update_proof proof (update_subproofs subproofs updates) | |
| 66 | and update_proof proof (proofs, []) = (proof :: proofs, []) | |
| 67 | | update_proof (Proof (fix, assms, steps)) (proofs, updates) = | |
| 68 | let val (steps, updates) = update_steps steps updates in | |
| 54712 | 69 | (Proof (fix, assms, steps) :: proofs, updates) | 
| 70 | end | |
| 52556 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 71 | in | 
| 55313 | 72 | fst (update_steps steps (rev updates)) | 
| 52556 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 73 | end | 
| 50672 
ab5b8b5c9cbe
added "obtain" to Isar proof construction data structure
 blanchet parents: 
50557diff
changeset | 74 | |
| 55283 
b90c06d54d38
when merging, don't try methods that failed or timed out for either of the steps for the combined step
 blanchet parents: 
55280diff
changeset | 75 | fun merge_methods preplay_data (l1, meths1) (l2, meths2) = | 
| 
b90c06d54d38
when merging, don't try methods that failed or timed out for either of the steps for the combined step
 blanchet parents: 
55280diff
changeset | 76 | let | 
| 55328 | 77 | fun is_hopeful l meth = | 
| 55283 
b90c06d54d38
when merging, don't try methods that failed or timed out for either of the steps for the combined step
 blanchet parents: 
55280diff
changeset | 78 | let val outcome = preplay_outcome_of_isar_step_for_method preplay_data l meth in | 
| 
b90c06d54d38
when merging, don't try methods that failed or timed out for either of the steps for the combined step
 blanchet parents: 
55280diff
changeset | 79 | not (Lazy.is_finished outcome) orelse | 
| 55328 | 80 | (case Lazy.force outcome of Played _ => true | Play_Timed_Out _ => true | _ => false) | 
| 55283 
b90c06d54d38
when merging, don't try methods that failed or timed out for either of the steps for the combined step
 blanchet parents: 
55280diff
changeset | 81 | end | 
| 55328 | 82 | |
| 83 | val (hopeful, hopeless) = | |
| 84 | meths2 @ subtract (op =) meths2 meths1 | |
| 85 | |> List.partition (is_hopeful l1 andf is_hopeful l2) | |
| 55283 
b90c06d54d38
when merging, don't try methods that failed or timed out for either of the steps for the combined step
 blanchet parents: 
55280diff
changeset | 86 | in | 
| 55328 | 87 | (hopeful @ hopeless, hopeless) | 
| 55283 
b90c06d54d38
when merging, don't try methods that failed or timed out for either of the steps for the combined step
 blanchet parents: 
55280diff
changeset | 88 | end | 
| 
b90c06d54d38
when merging, don't try methods that failed or timed out for either of the steps for the combined step
 blanchet parents: 
55280diff
changeset | 89 | |
| 55332 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 90 | fun merge_steps preplay_data (Prove ([], fix1, l1, _, subproofs1, (lfs1, gfs1), meths1, comment1)) | 
| 55329 
3c2dbd2e221f
more generous Isar proof compression -- try to remove failing steps
 blanchet parents: 
55328diff
changeset | 91 | (Prove (qs2, fix2, l2, t, subproofs2, (lfs2, gfs2), meths2, comment2)) = | 
| 55332 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 92 | let | 
| 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 93 | val (meths, hopeless) = merge_methods preplay_data (l1, meths1) (l2, meths2) | 
| 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 94 | val lfs = union (op =) lfs1 (remove (op =) l1 lfs2) | 
| 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 95 | val gfs = union (op =) gfs1 gfs2 | 
| 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 96 | in | 
| 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 97 | (Prove (qs2, union (op =) fix1 fix2, l2, t, subproofs1 @ subproofs2, (lfs, gfs), meths, | 
| 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 98 | comment1 ^ comment2), hopeless) | 
| 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 99 | end | 
| 50672 
ab5b8b5c9cbe
added "obtain" to Isar proof construction data structure
 blanchet parents: 
50557diff
changeset | 100 | |
| 55329 
3c2dbd2e221f
more generous Isar proof compression -- try to remove failing steps
 blanchet parents: 
55328diff
changeset | 101 | val merge_slack_time = seconds 0.005 | 
| 
3c2dbd2e221f
more generous Isar proof compression -- try to remove failing steps
 blanchet parents: 
55328diff
changeset | 102 | val merge_slack_factor = 1.5 | 
| 
3c2dbd2e221f
more generous Isar proof compression -- try to remove failing steps
 blanchet parents: 
55328diff
changeset | 103 | |
| 
3c2dbd2e221f
more generous Isar proof compression -- try to remove failing steps
 blanchet parents: 
55328diff
changeset | 104 | fun adjust_merge_timeout max time = | 
| 
3c2dbd2e221f
more generous Isar proof compression -- try to remove failing steps
 blanchet parents: 
55328diff
changeset | 105 | let val timeout = time_mult merge_slack_factor (Time.+ (merge_slack_time, time)) in | 
| 
3c2dbd2e221f
more generous Isar proof compression -- try to remove failing steps
 blanchet parents: 
55328diff
changeset | 106 | if Time.< (max, timeout) then max else timeout | 
| 
3c2dbd2e221f
more generous Isar proof compression -- try to remove failing steps
 blanchet parents: 
55328diff
changeset | 107 | end | 
| 
3c2dbd2e221f
more generous Isar proof compression -- try to remove failing steps
 blanchet parents: 
55328diff
changeset | 108 | |
| 53763 | 109 | val compress_degree = 2 | 
| 55271 
e78476a99c70
better time slack, to account for ultra-quick proof methods
 blanchet parents: 
55270diff
changeset | 110 | |
| 
e78476a99c70
better time slack, to account for ultra-quick proof methods
 blanchet parents: 
55270diff
changeset | 111 | (* Precondition: The proof must be labeled canonically. *) | 
| 55452 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55333diff
changeset | 112 | fun compress_isar_proof ctxt debug compress_isar preplay_timeout preplay_data proof = | 
| 55183 
17ec4a29ef71
renamed Sledgehammer options for symmetry between positive and negative versions
 blanchet parents: 
54828diff
changeset | 113 | if compress_isar <= 1.0 then | 
| 52556 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 114 | proof | 
| 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 115 | else | 
| 54712 | 116 | let | 
| 117 | val (compress_further, decrement_step_count) = | |
| 118 | let | |
| 55260 | 119 | val number_of_steps = add_isar_steps (steps_of_isar_proof proof) 0 | 
| 55183 
17ec4a29ef71
renamed Sledgehammer options for symmetry between positive and negative versions
 blanchet parents: 
54828diff
changeset | 120 | val target_number_of_steps = Real.round (Real.fromInt number_of_steps / compress_isar) | 
| 54712 | 121 | val delta = Unsynchronized.ref (number_of_steps - target_number_of_steps) | 
| 122 | in | |
| 123 | (fn () => !delta > 0, fn () => delta := !delta - 1) | |
| 124 | end | |
| 52556 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 125 | |
| 54712 | 126 | val (get_successors, replace_successor) = | 
| 127 | let | |
| 55299 | 128 | fun add_refs (Prove (_, _, l, _, _, (lfs, _), _, _)) = | 
| 129 | fold (fn key => Canonical_Label_Tab.cons_list (key, l)) lfs | |
| 130 | | add_refs _ = I | |
| 52556 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 131 | |
| 54712 | 132 | val tab = | 
| 55212 | 133 | Canonical_Label_Tab.empty | 
| 55260 | 134 | |> fold_isar_steps add_refs (steps_of_isar_proof proof) | 
| 54712 | 135 | (* "rev" should have the same effect as "sort canonical_label_ord" *) | 
| 55212 | 136 | |> Canonical_Label_Tab.map (K rev) | 
| 54712 | 137 | |> Unsynchronized.ref | 
| 51260 | 138 | |
| 55212 | 139 | fun get_successors l = Canonical_Label_Tab.lookup_list (!tab) l | 
| 140 | fun set_successors l refs = tab := Canonical_Label_Tab.update (l, refs) (!tab) | |
| 54712 | 141 | fun replace_successor old new dest = | 
| 142 | get_successors dest | |
| 143 | |> Ord_List.remove canonical_label_ord old | |
| 144 | |> Ord_List.union canonical_label_ord new | |
| 145 | |> set_successors dest | |
| 146 | in | |
| 147 | (get_successors, replace_successor) | |
| 148 | end | |
| 52556 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 149 | |
| 55333 
fa079fd40db8
more exception cleanup + more liberal compressions of steps that timed out
 blanchet parents: 
55332diff
changeset | 150 | fun reference_time l = | 
| 
fa079fd40db8
more exception cleanup + more liberal compressions of steps that timed out
 blanchet parents: 
55332diff
changeset | 151 | (case forced_intermediate_preplay_outcome_of_isar_step (!preplay_data) l of | 
| 
fa079fd40db8
more exception cleanup + more liberal compressions of steps that timed out
 blanchet parents: 
55332diff
changeset | 152 | Played time => time | 
| 
fa079fd40db8
more exception cleanup + more liberal compressions of steps that timed out
 blanchet parents: 
55332diff
changeset | 153 | | _ => preplay_timeout) | 
| 
fa079fd40db8
more exception cleanup + more liberal compressions of steps that timed out
 blanchet parents: 
55332diff
changeset | 154 | |
| 55243 | 155 | (* elimination of trivial, one-step subproofs *) | 
| 55330 
547d23e2abf7
corrected wrong 'meth :: _' pattern maching + tuned code
 blanchet parents: 
55329diff
changeset | 156 | fun elim_one_subproof time (step as Prove (qs, fix, l, t, _, (lfs, gfs), meths, comment)) subs | 
| 55309 | 157 | nontriv_subs = | 
| 54712 | 158 | if null subs orelse not (compress_further ()) then | 
| 55330 
547d23e2abf7
corrected wrong 'meth :: _' pattern maching + tuned code
 blanchet parents: 
55329diff
changeset | 159 | Prove (qs, fix, l, t, List.revAppend (nontriv_subs, subs), (lfs, gfs), meths, comment) | 
| 54712 | 160 | else | 
| 161 | (case subs of | |
| 55333 
fa079fd40db8
more exception cleanup + more liberal compressions of steps that timed out
 blanchet parents: 
55332diff
changeset | 162 | (sub as Proof (_, assms, [Prove (_, _, l', _, [], (lfs', gfs'), meths', _)])) :: subs => | 
| 
fa079fd40db8
more exception cleanup + more liberal compressions of steps that timed out
 blanchet parents: 
55332diff
changeset | 163 | let | 
| 52556 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 164 | (* merge steps *) | 
| 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 165 | val subs'' = subs @ nontriv_subs | 
| 55223 
3c593bad6b31
generalized preplaying infrastructure to store various results for various methods
 blanchet parents: 
55221diff
changeset | 166 | val lfs'' = union (op =) lfs (subtract (op =) (map fst assms) lfs') | 
| 52556 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 167 | val gfs'' = union (op =) gfs' gfs | 
| 55328 | 168 | val (meths'' as _ :: _, hopeless) = | 
| 169 | merge_methods (!preplay_data) (l', meths') (l, meths) | |
| 55299 | 170 | val step'' = Prove (qs, fix, l, t, subs'', (lfs'', gfs''), meths'', comment) | 
| 52556 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 171 | |
| 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 172 | (* check if the modified step can be preplayed fast enough *) | 
| 55333 
fa079fd40db8
more exception cleanup + more liberal compressions of steps that timed out
 blanchet parents: 
55332diff
changeset | 173 | val timeout = adjust_merge_timeout preplay_timeout (Time.+ (time, reference_time l')) | 
| 52556 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 174 | in | 
| 55452 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55333diff
changeset | 175 | (case preplay_isar_step ctxt debug timeout hopeless step'' of | 
| 55333 
fa079fd40db8
more exception cleanup + more liberal compressions of steps that timed out
 blanchet parents: 
55332diff
changeset | 176 | meths_outcomes as (_, Played time'') :: _ => | 
| 
fa079fd40db8
more exception cleanup + more liberal compressions of steps that timed out
 blanchet parents: 
55332diff
changeset | 177 | (* l' successfully eliminated *) | 
| 
fa079fd40db8
more exception cleanup + more liberal compressions of steps that timed out
 blanchet parents: 
55332diff
changeset | 178 | (decrement_step_count (); | 
| 55452 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55333diff
changeset | 179 | set_preplay_outcomes_of_isar_step ctxt debug time'' preplay_data step'' | 
| 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55333diff
changeset | 180 | meths_outcomes; | 
| 55333 
fa079fd40db8
more exception cleanup + more liberal compressions of steps that timed out
 blanchet parents: 
55332diff
changeset | 181 | map (replace_successor l' [l]) lfs'; | 
| 
fa079fd40db8
more exception cleanup + more liberal compressions of steps that timed out
 blanchet parents: 
55332diff
changeset | 182 | elim_one_subproof time'' step'' subs nontriv_subs) | 
| 
fa079fd40db8
more exception cleanup + more liberal compressions of steps that timed out
 blanchet parents: 
55332diff
changeset | 183 | | _ => elim_one_subproof time step subs (sub :: nontriv_subs)) | 
| 52556 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 184 | end | 
| 55333 
fa079fd40db8
more exception cleanup + more liberal compressions of steps that timed out
 blanchet parents: 
55332diff
changeset | 185 | | sub :: subs => elim_one_subproof time step subs (sub :: nontriv_subs)) | 
| 52556 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 186 | |
| 55333 
fa079fd40db8
more exception cleanup + more liberal compressions of steps that timed out
 blanchet parents: 
55332diff
changeset | 187 | fun elim_subproofs (step as Prove (_, _, l, _, subproofs, _, _, _)) = | 
| 
fa079fd40db8
more exception cleanup + more liberal compressions of steps that timed out
 blanchet parents: 
55332diff
changeset | 188 | if exists (null o tl o steps_of_isar_proof) subproofs then | 
| 
fa079fd40db8
more exception cleanup + more liberal compressions of steps that timed out
 blanchet parents: 
55332diff
changeset | 189 | elim_one_subproof (reference_time l) step subproofs [] | 
| 
fa079fd40db8
more exception cleanup + more liberal compressions of steps that timed out
 blanchet parents: 
55332diff
changeset | 190 | else | 
| 
fa079fd40db8
more exception cleanup + more liberal compressions of steps that timed out
 blanchet parents: 
55332diff
changeset | 191 | step | 
| 55299 | 192 | | elim_subproofs step = step | 
| 52556 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 193 | |
| 54712 | 194 | fun compress_top_level steps = | 
| 195 | let | |
| 55331 | 196 | fun cand_key (l, t_size) = (length (get_successors l), t_size) | 
| 197 | val cand_ord = prod_ord int_ord (int_ord o swap) o pairself cand_key | |
| 52556 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 198 | |
| 55298 | 199 | fun pop_next_candidate [] = (NONE, []) | 
| 200 | | pop_next_candidate (cands as (cand :: cands')) = | |
| 55331 | 201 | fold (fn x => fn y => if cand_ord (x, y) = LESS then x else y) cands' cand | 
| 55325 | 202 | |> (fn best => (SOME best, remove (op =) best cands)) | 
| 52556 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 203 | |
| 55332 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 204 | fun try_eliminate i l labels steps = | 
| 54712 | 205 | let | 
| 55330 
547d23e2abf7
corrected wrong 'meth :: _' pattern maching + tuned code
 blanchet parents: 
55329diff
changeset | 206 | val (steps_before, (cand as Prove (_, _, _, _, _, (lfs, _), _, _)) :: steps_after) = | 
| 55332 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 207 | chop i steps | 
| 55330 
547d23e2abf7
corrected wrong 'meth :: _' pattern maching + tuned code
 blanchet parents: 
55329diff
changeset | 208 | val succs = collect_successors steps_after labels | 
| 55332 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 209 | val (succs', hopelesses) = split_list (map (merge_steps (!preplay_data) cand) succs) | 
| 54712 | 210 | in | 
| 55332 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 211 | (case try (map ((fn Played time => time) o | 
| 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 212 | forced_intermediate_preplay_outcome_of_isar_step (!preplay_data))) labels of | 
| 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 213 | NONE => steps | 
| 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 214 | | SOME times0 => | 
| 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 215 | let | 
| 55333 
fa079fd40db8
more exception cleanup + more liberal compressions of steps that timed out
 blanchet parents: 
55332diff
changeset | 216 | val time_slice = time_mult (1.0 / Real.fromInt (length labels)) (reference_time l) | 
| 55332 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 217 | val timeouts = | 
| 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 218 | map (adjust_merge_timeout preplay_timeout o curry Time.+ time_slice) times0 | 
| 55452 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55333diff
changeset | 219 | val meths_outcomess = | 
| 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55333diff
changeset | 220 | map3 (preplay_isar_step ctxt debug) timeouts hopelesses succs' | 
| 55332 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 221 | in | 
| 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 222 | (case try (map (fn (_, Played time) :: _ => time)) meths_outcomess of | 
| 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 223 | NONE => steps | 
| 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 224 | | SOME times => | 
| 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 225 | (* candidate successfully eliminated *) | 
| 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 226 | (decrement_step_count (); | 
| 55452 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55333diff
changeset | 227 | map3 (fn time => | 
| 
29ec8680e61f
avoid changing the state's context -- this results in transfer problems later with SMT, and hence preplay tactic failures
 blanchet parents: 
55333diff
changeset | 228 | set_preplay_outcomes_of_isar_step ctxt debug time preplay_data) | 
| 55332 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 229 | times succs' meths_outcomess; | 
| 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 230 | map (replace_successor l labels) lfs; | 
| 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 231 | steps_before @ update_steps succs' steps_after)) | 
| 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 232 | end) | 
| 54712 | 233 | end | 
| 52556 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 234 | |
| 54712 | 235 | fun compression_loop candidates steps = | 
| 236 | if not (compress_further ()) then | |
| 237 | steps | |
| 238 | else | |
| 55298 | 239 | (case pop_next_candidate candidates of | 
| 54712 | 240 | (NONE, _) => steps (* no more candidates for elimination *) | 
| 55332 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 241 | | (SOME (l, _), candidates') => | 
| 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 242 | (case find_index (curry (op =) (SOME l) o label_of_isar_step) steps of | 
| 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 243 | ~1 => steps | 
| 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 244 | | i => | 
| 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 245 | let val successors = get_successors l in | 
| 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 246 | if length successors > compress_degree then steps | 
| 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 247 | else compression_loop candidates' (try_eliminate i l successors steps) | 
| 
803a7400cc58
tuned code to avoid noncanonical (and risky) exception handling
 blanchet parents: 
55331diff
changeset | 248 | end)) | 
| 55331 | 249 | |
| 250 | fun add_cand (Prove (_, _, l, t, _, _, _, _)) = cons (l, size_of_term t) | |
| 251 | | add_cand _ = I | |
| 252 | ||
| 253 | (* the very last step is not a candidate *) | |
| 254 | val candidates = fold add_cand (fst (split_last steps)) [] | |
| 54712 | 255 | in | 
| 256 | compression_loop candidates steps | |
| 257 | end | |
| 52556 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 258 | |
| 55268 | 259 | (* Proofs are compressed bottom-up, beginning with the innermost subproofs. On the innermost | 
| 260 | proof level, the proof steps have no subproofs. In the best case, these steps can be merged | |
| 261 | into just one step, resulting in a trivial subproof. Going one level up, trivial subproofs | |
| 262 | can be eliminated. In the best case, this once again leads to a proof whose proof steps do | |
| 263 | not have subproofs. Applying this approach recursively will result in a flat proof in the | |
| 264 | best cast. *) | |
| 55263 | 265 | fun compress_proof (proof as (Proof (fix, assms, steps))) = | 
| 266 | if compress_further () then Proof (fix, assms, compress_steps steps) else proof | |
| 267 | and compress_steps steps = | |
| 54712 | 268 | (* bottom-up: compress innermost proofs first *) | 
| 55263 | 269 | steps |> map (fn step => step |> compress_further () ? compress_sub_levels) | 
| 54712 | 270 | |> compress_further () ? compress_top_level | 
| 55329 
3c2dbd2e221f
more generous Isar proof compression -- try to remove failing steps
 blanchet parents: 
55328diff
changeset | 271 | and compress_sub_levels (Prove (qs, xs, l, t, subproofs, facts, meths, comment)) = | 
| 52556 
c8357085217c
completely rewrote SH compress; added two parameters for experimentation/fine grained control
 smolkas parents: 
52454diff
changeset | 272 | (* compress subproofs *) | 
| 55299 | 273 | Prove (qs, xs, l, t, map compress_proof subproofs, facts, meths, comment) | 
| 54712 | 274 | (* eliminate trivial subproofs *) | 
| 275 | |> compress_further () ? elim_subproofs | |
| 55329 
3c2dbd2e221f
more generous Isar proof compression -- try to remove failing steps
 blanchet parents: 
55328diff
changeset | 276 | | compress_sub_levels step = step | 
| 54712 | 277 | in | 
| 55263 | 278 | compress_proof proof | 
| 54712 | 279 | end | 
| 50259 | 280 | |
| 54504 | 281 | end; |