src/HOL/Tools/Sledgehammer/sledgehammer_isar_compress.ML
author blanchet
Tue, 04 Feb 2014 23:11:18 +0100
changeset 55323 253a029335a2
parent 55313 cddd94fb0e8d
child 55325 462ffd3b7065
permissions -rw-r--r--
split 'linarith' and 'presburger' (to avoid annoying warnings + to speed up reconstruction when 'presburger' is needed)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
55202
824c48a539c9 renamed many Sledgehammer ML files to clarify structure
blanchet
parents: 55183
diff changeset
     1
(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_isar_compress.ML
54813
blanchet
parents: 54752
diff changeset
     2
    Author:     Steffen Juilf Smolka, TU Muenchen
50263
0b430064296a added comments to new source files
smolkas
parents: 50261
diff changeset
     3
    Author:     Jasmin Blanchette, TU Muenchen
0b430064296a added comments to new source files
smolkas
parents: 50261
diff changeset
     4
55202
824c48a539c9 renamed many Sledgehammer ML files to clarify structure
blanchet
parents: 55183
diff changeset
     5
Compression of Isar proofs by merging steps.
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
     6
Only proof steps using the same proof method are merged.
50263
0b430064296a added comments to new source files
smolkas
parents: 50261
diff changeset
     7
*)
0b430064296a added comments to new source files
smolkas
parents: 50261
diff changeset
     8
55202
824c48a539c9 renamed many Sledgehammer ML files to clarify structure
blanchet
parents: 55183
diff changeset
     9
signature SLEDGEHAMMER_ISAR_COMPRESS =
50259
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    10
sig
55202
824c48a539c9 renamed many Sledgehammer ML files to clarify structure
blanchet
parents: 55183
diff changeset
    11
  type isar_proof = Sledgehammer_Isar_Proof.isar_proof
55213
dcb36a2540bc tuned ML function names
blanchet
parents: 55212
diff changeset
    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: 52454
diff changeset
    13
55260
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
    14
  val compress_isar_proof : Proof.context -> real -> isar_preplay_data Unsynchronized.ref ->
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
    15
    isar_proof -> isar_proof
54504
blanchet
parents: 54503
diff changeset
    16
end;
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
    17
55202
824c48a539c9 renamed many Sledgehammer ML files to clarify structure
blanchet
parents: 55183
diff changeset
    18
structure Sledgehammer_Isar_Compress : SLEDGEHAMMER_ISAR_COMPRESS =
50259
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    19
struct
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    20
50265
9eafa567e061 made use of sledgehammer_util
smolkas
parents: 50264
diff changeset
    21
open Sledgehammer_Util
55287
ffa306239316 renamed ML file
blanchet
parents: 55283
diff changeset
    22
open Sledgehammer_Proof_Methods
55202
824c48a539c9 renamed many Sledgehammer ML files to clarify structure
blanchet
parents: 55183
diff changeset
    23
open Sledgehammer_Isar_Proof
824c48a539c9 renamed many Sledgehammer ML files to clarify structure
blanchet
parents: 55183
diff changeset
    24
open Sledgehammer_Isar_Preplay
50259
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    25
55212
blanchet
parents: 55202
diff changeset
    26
val dummy_isar_step = Let (Term.dummy, Term.dummy)
blanchet
parents: 55202
diff changeset
    27
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
    28
(* traverses steps in post-order and collects the steps with the given labels *)
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
    29
fun collect_successors steps lbls =
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
    30
  let
55263
blanchet
parents: 55260
diff changeset
    31
    fun collect_steps _ ([], accu) = ([], accu)
blanchet
parents: 55260
diff changeset
    32
      | collect_steps [] accum = accum
blanchet
parents: 55260
diff changeset
    33
      | collect_steps (step :: steps) accum = collect_steps steps (collect_step step accum)
blanchet
parents: 55260
diff changeset
    34
    and collect_step (Let _) x = x
55299
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
    35
      | collect_step (step as Prove (_, _, l, _, subproofs, _, _, _)) x =
55263
blanchet
parents: 55260
diff changeset
    36
        (case collect_subproofs subproofs x of
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
    37
          ([], accu) => ([], accu)
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
    38
        | accum as (l' :: lbls', accu) => if l = l' then (lbls', step :: accu) else accum)
55263
blanchet
parents: 55260
diff changeset
    39
    and collect_subproofs [] x = x
blanchet
parents: 55260
diff changeset
    40
      | collect_subproofs (proof :: subproofs) x =
blanchet
parents: 55260
diff changeset
    41
        (case collect_steps (steps_of_isar_proof proof) x of
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
    42
          accum as ([], _) => accum
55263
blanchet
parents: 55260
diff changeset
    43
        | accum => collect_subproofs subproofs accum)
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
    44
  in
55313
blanchet
parents: 55312
diff changeset
    45
    rev (snd (collect_steps steps (lbls, [])))
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
    46
  end
50259
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    47
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
    48
(* traverses steps in reverse post-order and inserts the given updates *)
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
    49
fun update_steps steps updates =
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
    50
  let
55263
blanchet
parents: 55260
diff changeset
    51
    fun update_steps [] updates = ([], updates)
blanchet
parents: 55260
diff changeset
    52
      | update_steps steps [] = (steps, [])
blanchet
parents: 55260
diff changeset
    53
      | update_steps (step :: steps) updates = update_step step (update_steps steps updates)
blanchet
parents: 55260
diff changeset
    54
    and update_step step (steps, []) = (step :: steps, [])
blanchet
parents: 55260
diff changeset
    55
      | update_step (step as Let _) (steps, updates) = (step :: steps, updates)
55299
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
    56
      | update_step (Prove (qs, xs, l, t, subproofs, facts, meths, comment))
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
    57
          (steps,
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
    58
           updates as Prove (qs', xs', l', t', subproofs', facts', meths', comment') :: updates') =
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
    59
        (if l = l' then
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
    60
           update_subproofs subproofs' updates'
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
    61
           |>> (fn subproofs' => Prove (qs', xs', l', t', subproofs', facts', meths', comment'))
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
    62
         else
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
    63
           update_subproofs subproofs updates
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
    64
           |>> (fn subproofs => Prove (qs, xs, l, t, subproofs, facts, meths, comment)))
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
    65
        |>> (fn step => step :: steps)
55263
blanchet
parents: 55260
diff changeset
    66
    and update_subproofs [] updates = ([], updates)
blanchet
parents: 55260
diff changeset
    67
      | update_subproofs steps [] = (steps, [])
blanchet
parents: 55260
diff changeset
    68
      | update_subproofs (proof :: subproofs) updates =
blanchet
parents: 55260
diff changeset
    69
        update_proof proof (update_subproofs subproofs updates)
blanchet
parents: 55260
diff changeset
    70
    and update_proof proof (proofs, []) = (proof :: proofs, [])
blanchet
parents: 55260
diff changeset
    71
      | update_proof (Proof (fix, assms, steps)) (proofs, updates) =
blanchet
parents: 55260
diff changeset
    72
        let val (steps, updates) = update_steps steps updates in
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
    73
          (Proof (fix, assms, steps) :: proofs, updates)
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
    74
        end
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
    75
  in
55313
blanchet
parents: 55312
diff changeset
    76
    fst (update_steps steps (rev updates))
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
    77
  end
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
    78
55283
b90c06d54d38 when merging, don't try methods that failed or timed out for either of the steps for the combined step
blanchet
parents: 55280
diff changeset
    79
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: 55280
diff changeset
    80
  let
b90c06d54d38 when merging, don't try methods that failed or timed out for either of the steps for the combined step
blanchet
parents: 55280
diff changeset
    81
    fun is_method_hopeful l meth =
b90c06d54d38 when merging, don't try methods that failed or timed out for either of the steps for the combined step
blanchet
parents: 55280
diff changeset
    82
      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: 55280
diff changeset
    83
        not (Lazy.is_finished outcome) orelse
b90c06d54d38 when merging, don't try methods that failed or timed out for either of the steps for the combined step
blanchet
parents: 55280
diff changeset
    84
        (case Lazy.force outcome of Played _ => true | _ => false)
b90c06d54d38 when merging, don't try methods that failed or timed out for either of the steps for the combined step
blanchet
parents: 55280
diff changeset
    85
      end
b90c06d54d38 when merging, don't try methods that failed or timed out for either of the steps for the combined step
blanchet
parents: 55280
diff changeset
    86
  in
55323
253a029335a2 split 'linarith' and 'presburger' (to avoid annoying warnings + to speed up reconstruction when 'presburger' is needed)
blanchet
parents: 55313
diff changeset
    87
    meths2 @ subtract (op =) meths2 meths1
55312
e7029ee73a97 more liberal step merging
blanchet
parents: 55309
diff changeset
    88
    |> filter (is_method_hopeful l1 andf is_method_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: 55280
diff changeset
    89
  end
b90c06d54d38 when merging, don't try methods that failed or timed out for either of the steps for the combined step
blanchet
parents: 55280
diff changeset
    90
55299
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
    91
fun try_merge preplay_data (Prove (_, [], l1, _, subproofs1, (lfs1, gfs1), meths1, comment1))
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
    92
      (Prove (qs2, fix, l2, t, subproofs2, (lfs2, gfs2), meths2, comment2)) =
55283
b90c06d54d38 when merging, don't try methods that failed or timed out for either of the steps for the combined step
blanchet
parents: 55280
diff changeset
    93
    (case merge_methods preplay_data (l1, meths1) (l2, meths2) of
55247
4aa3e1c6222c take intersection rather than union of methods when merging steps -- more efficient and natural
blanchet
parents: 55246
diff changeset
    94
      [] => NONE
4aa3e1c6222c take intersection rather than union of methods when merging steps -- more efficient and natural
blanchet
parents: 55246
diff changeset
    95
    | meths =>
4aa3e1c6222c take intersection rather than union of methods when merging steps -- more efficient and natural
blanchet
parents: 55246
diff changeset
    96
      let
4aa3e1c6222c take intersection rather than union of methods when merging steps -- more efficient and natural
blanchet
parents: 55246
diff changeset
    97
        val lfs = union (op =) lfs1 (remove (op =) l1 lfs2)
4aa3e1c6222c take intersection rather than union of methods when merging steps -- more efficient and natural
blanchet
parents: 55246
diff changeset
    98
        val gfs = union (op =) gfs1 gfs2
4aa3e1c6222c take intersection rather than union of methods when merging steps -- more efficient and natural
blanchet
parents: 55246
diff changeset
    99
      in
55299
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
   100
        SOME (Prove (qs2, fix, l2, t, subproofs1 @ subproofs2, (lfs, gfs), meths,
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
   101
          comment1 ^ comment2))
55247
4aa3e1c6222c take intersection rather than union of methods when merging steps -- more efficient and natural
blanchet
parents: 55246
diff changeset
   102
      end)
55283
b90c06d54d38 when merging, don't try methods that failed or timed out for either of the steps for the combined step
blanchet
parents: 55280
diff changeset
   103
  | try_merge _ _ _ = NONE
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   104
53763
70d370743dc6 hardcoded obscure option
blanchet
parents: 53762
diff changeset
   105
val compress_degree = 2
55271
e78476a99c70 better time slack, to account for ultra-quick proof methods
blanchet
parents: 55270
diff changeset
   106
val merge_timeout_slack_time = seconds 0.005
55278
73372494fd80 more flexible compression, choosing whichever proof method works
blanchet
parents: 55272
diff changeset
   107
val merge_timeout_slack_factor = 1.25
53762
06510d01a07b hard-coded an obscure option
blanchet
parents: 52692
diff changeset
   108
55271
e78476a99c70 better time slack, to account for ultra-quick proof methods
blanchet
parents: 55270
diff changeset
   109
fun slackify_merge_timeout time =
e78476a99c70 better time slack, to account for ultra-quick proof methods
blanchet
parents: 55270
diff changeset
   110
  time_mult merge_timeout_slack_factor (Time.+ (merge_timeout_slack_time, time))
e78476a99c70 better time slack, to account for ultra-quick proof methods
blanchet
parents: 55270
diff changeset
   111
e78476a99c70 better time slack, to account for ultra-quick proof methods
blanchet
parents: 55270
diff changeset
   112
(* Precondition: The proof must be labeled canonically. *)
55260
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
   113
fun compress_isar_proof ctxt compress_isar preplay_data proof =
55183
17ec4a29ef71 renamed Sledgehammer options for symmetry between positive and negative versions
blanchet
parents: 54828
diff changeset
   114
  if compress_isar <= 1.0 then
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   115
    proof
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   116
  else
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   117
    let
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   118
      val (compress_further, decrement_step_count) =
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   119
        let
55260
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
   120
          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: 54828
diff changeset
   121
          val target_number_of_steps = Real.round (Real.fromInt number_of_steps / compress_isar)
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   122
          val delta = Unsynchronized.ref (number_of_steps - target_number_of_steps)
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   123
        in
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   124
          (fn () => !delta > 0, fn () => delta := !delta - 1)
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   125
        end
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   126
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   127
      val (get_successors, replace_successor) =
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   128
        let
55299
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
   129
          fun add_refs (Prove (_, _, l, _, _, (lfs, _), _, _)) =
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
   130
              fold (fn key => Canonical_Label_Tab.cons_list (key, l)) lfs
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
   131
            | add_refs _ = I
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   132
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   133
          val tab =
55212
blanchet
parents: 55202
diff changeset
   134
            Canonical_Label_Tab.empty
55260
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
   135
            |> fold_isar_steps add_refs (steps_of_isar_proof proof)
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   136
            (* "rev" should have the same effect as "sort canonical_label_ord" *)
55212
blanchet
parents: 55202
diff changeset
   137
            |> Canonical_Label_Tab.map (K rev)
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   138
            |> Unsynchronized.ref
51260
61bc5a3bef09 tuned agressiveness of isar compression
smolkas
parents: 51179
diff changeset
   139
55212
blanchet
parents: 55202
diff changeset
   140
          fun get_successors l = Canonical_Label_Tab.lookup_list (!tab) l
blanchet
parents: 55202
diff changeset
   141
          fun set_successors l refs = tab := Canonical_Label_Tab.update (l, refs) (!tab)
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   142
          fun replace_successor old new dest =
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   143
            get_successors dest
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   144
            |> Ord_List.remove canonical_label_ord old
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   145
            |> Ord_List.union canonical_label_ord new
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   146
            |> set_successors dest
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   147
        in
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   148
          (get_successors, replace_successor)
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   149
        end
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   150
55243
66709d41601e reset timing information after changes
blanchet
parents: 55223
diff changeset
   151
      (* elimination of trivial, one-step subproofs *)
55309
455a7f9924df don't lose additional outcomes
blanchet
parents: 55307
diff changeset
   152
      fun elim_one_subproof time meths_outcomes qs fix l t lfs gfs (meths as meth :: _) comment subs
455a7f9924df don't lose additional outcomes
blanchet
parents: 55307
diff changeset
   153
          nontriv_subs =
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   154
        if null subs orelse not (compress_further ()) then
55264
43473497fb65 centralize preplaying
blanchet
parents: 55263
diff changeset
   155
          let
43473497fb65 centralize preplaying
blanchet
parents: 55263
diff changeset
   156
            val subproofs = List.revAppend (nontriv_subs, subs)
55299
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
   157
            val step = Prove (qs, fix, l, t, subproofs, (lfs, gfs), meths, comment)
55264
43473497fb65 centralize preplaying
blanchet
parents: 55263
diff changeset
   158
          in
55309
455a7f9924df don't lose additional outcomes
blanchet
parents: 55307
diff changeset
   159
            set_preplay_outcomes_of_isar_step ctxt time preplay_data step
455a7f9924df don't lose additional outcomes
blanchet
parents: 55307
diff changeset
   160
              ((meth, Played time) :: meths_outcomes);
55264
43473497fb65 centralize preplaying
blanchet
parents: 55263
diff changeset
   161
            step
43473497fb65 centralize preplaying
blanchet
parents: 55263
diff changeset
   162
          end
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   163
        else
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   164
          (case subs of
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   165
            (sub as Proof (_, assms, sub_steps)) :: subs =>
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   166
            (let
55223
3c593bad6b31 generalized preplaying infrastructure to store various results for various methods
blanchet
parents: 55221
diff changeset
   167
              (* trivial subproofs have exactly one "Prove" step *)
55299
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
   168
              val [Prove (_, [], l', _, [], (lfs', gfs'), meths', _)] = sub_steps
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   169
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   170
              (* only touch proofs that can be preplayed sucessfully *)
55272
blanchet
parents: 55271
diff changeset
   171
              val Played time' = forced_intermediate_preplay_outcome_of_isar_step (!preplay_data) l'
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   172
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   173
              (* merge steps *)
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   174
              val subs'' = subs @ nontriv_subs
55223
3c593bad6b31 generalized preplaying infrastructure to store various results for various methods
blanchet
parents: 55221
diff changeset
   175
              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: 52454
diff changeset
   176
              val gfs'' = union (op =) gfs' gfs
55283
b90c06d54d38 when merging, don't try methods that failed or timed out for either of the steps for the combined step
blanchet
parents: 55280
diff changeset
   177
              val meths'' as _ :: _ = merge_methods (!preplay_data) (l', meths') (l, meths)
55299
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
   178
              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: 52454
diff changeset
   179
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   180
              (* check if the modified step can be preplayed fast enough *)
55271
e78476a99c70 better time slack, to account for ultra-quick proof methods
blanchet
parents: 55270
diff changeset
   181
              val timeout = slackify_merge_timeout (Time.+ (time, time'))
55309
455a7f9924df don't lose additional outcomes
blanchet
parents: 55307
diff changeset
   182
              val (_, Played time'') :: meths_outcomes = preplay_isar_step ctxt timeout step''
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   183
            in
52626
79a4e7f8d758 cleaner preplay interface
smolkas
parents: 52614
diff changeset
   184
              decrement_step_count (); (* l' successfully eliminated! *)
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   185
              map (replace_successor l' [l]) lfs';
55309
455a7f9924df don't lose additional outcomes
blanchet
parents: 55307
diff changeset
   186
              elim_one_subproof time'' meths_outcomes qs fix l t lfs'' gfs'' meths comment subs
455a7f9924df don't lose additional outcomes
blanchet
parents: 55307
diff changeset
   187
                nontriv_subs
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
            handle Bind =>
55309
455a7f9924df don't lose additional outcomes
blanchet
parents: 55307
diff changeset
   190
              elim_one_subproof time [] qs fix l t lfs gfs meths comment subs (sub :: nontriv_subs))
55264
43473497fb65 centralize preplaying
blanchet
parents: 55263
diff changeset
   191
          | _ => raise Fail "Sledgehammer_Isar_Compress: elim_one_subproof")
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   192
55299
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
   193
      fun elim_subproofs (step as Prove (qs, fix, l, t, subproofs, (lfs, gfs), meths, comment)) =
54813
blanchet
parents: 54752
diff changeset
   194
          if subproofs = [] then
blanchet
parents: 54752
diff changeset
   195
            step
blanchet
parents: 54752
diff changeset
   196
          else
55272
blanchet
parents: 55271
diff changeset
   197
            (case forced_intermediate_preplay_outcome_of_isar_step (!preplay_data) l of
55309
455a7f9924df don't lose additional outcomes
blanchet
parents: 55307
diff changeset
   198
              Played time => elim_one_subproof time [] qs fix l t lfs gfs meths comment subproofs []
54826
79745ba60a5a data structure rationalization
blanchet
parents: 54813
diff changeset
   199
            | _ => step)
55299
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
   200
        | elim_subproofs step = step
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   201
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   202
      fun compress_top_level steps =
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   203
        let
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   204
          (* (#successors, (size_of_term t, position)) *)
55223
3c593bad6b31 generalized preplaying infrastructure to store various results for various methods
blanchet
parents: 55221
diff changeset
   205
          fun cand_key (i, l, t_size) = (length (get_successors l), (t_size, i))
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   206
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   207
          val compression_ord =
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   208
            prod_ord int_ord (prod_ord (int_ord #> rev_order) int_ord)
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   209
            #> rev_order
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   210
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   211
          val cand_ord = pairself cand_key #> compression_ord
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   212
55298
53569ca831f4 allow merging of steps with subproofs
blanchet
parents: 55287
diff changeset
   213
          fun pop_next_candidate [] = (NONE, [])
53569ca831f4 allow merging of steps with subproofs
blanchet
parents: 55287
diff changeset
   214
            | pop_next_candidate (cands as (cand :: cands')) =
55212
blanchet
parents: 55202
diff changeset
   215
              let
blanchet
parents: 55202
diff changeset
   216
                val best as (i, _, _) =
blanchet
parents: 55202
diff changeset
   217
                  fold (fn x => fn y => if cand_ord (x, y) = GREATER then x else y) cands' cand
blanchet
parents: 55202
diff changeset
   218
              in (SOME best, filter_out (fn (j, _, _) => j = i) cands) end
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   219
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   220
          val candidates =
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   221
            let
55299
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
   222
              fun add_cand (i, Prove (_, _, l, t, _, _, _, _)) = cons (i, l, size_of_term t)
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
   223
                | add_cand _ = I
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   224
            in
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   225
              (steps
55223
3c593bad6b31 generalized preplaying infrastructure to store various results for various methods
blanchet
parents: 55221
diff changeset
   226
               |> split_last |> fst (* keep last step *)
3c593bad6b31 generalized preplaying infrastructure to store various results for various methods
blanchet
parents: 55221
diff changeset
   227
               |> fold_index add_cand) []
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   228
            end
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   229
55264
43473497fb65 centralize preplaying
blanchet
parents: 55263
diff changeset
   230
          fun try_eliminate (i, l, _) labels steps =
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   231
            let
55299
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
   232
              val ((cand as Prove (_, _, _, _, _, (lfs, _), _, _)) :: steps') = drop i steps
55223
3c593bad6b31 generalized preplaying infrastructure to store various results for various methods
blanchet
parents: 55221
diff changeset
   233
55264
43473497fb65 centralize preplaying
blanchet
parents: 55263
diff changeset
   234
              val succs = collect_successors steps' labels
55223
3c593bad6b31 generalized preplaying infrastructure to store various results for various methods
blanchet
parents: 55221
diff changeset
   235
55307
blanchet
parents: 55299
diff changeset
   236
              (* only touch steps that can be preplayed successfully; FIXME: more generous *)
55272
blanchet
parents: 55271
diff changeset
   237
              val Played time = forced_intermediate_preplay_outcome_of_isar_step (!preplay_data) l
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   238
55283
b90c06d54d38 when merging, don't try methods that failed or timed out for either of the steps for the combined step
blanchet
parents: 55280
diff changeset
   239
              val succs' = map (try_merge (!preplay_data) cand #> the) succs
55223
3c593bad6b31 generalized preplaying infrastructure to store various results for various methods
blanchet
parents: 55221
diff changeset
   240
55307
blanchet
parents: 55299
diff changeset
   241
              (* FIXME: more generous *)
55269
aae87746f412 more thorough, hybrid compression
blanchet
parents: 55268
diff changeset
   242
              val times0 = map ((fn Played time => time) o
55272
blanchet
parents: 55271
diff changeset
   243
                forced_intermediate_preplay_outcome_of_isar_step (!preplay_data)) labels
55264
43473497fb65 centralize preplaying
blanchet
parents: 55263
diff changeset
   244
              val time_slice = time_mult (1.0 / (Real.fromInt (length labels))) time
55271
e78476a99c70 better time slack, to account for ultra-quick proof methods
blanchet
parents: 55270
diff changeset
   245
              val timeouts = map (curry Time.+ time_slice #> slackify_merge_timeout) times0
55307
blanchet
parents: 55299
diff changeset
   246
              (* FIXME: "preplay_timeout" should be an ultimate maximum *)
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   247
55278
73372494fd80 more flexible compression, choosing whichever proof method works
blanchet
parents: 55272
diff changeset
   248
              val meths_outcomess = map2 (preplay_isar_step ctxt) timeouts succs'
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   249
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   250
              (* ensure none of the modified successors timed out *)
55278
73372494fd80 more flexible compression, choosing whichever proof method works
blanchet
parents: 55272
diff changeset
   251
              val times = map (fn (_, Played time) :: _ => time) meths_outcomess
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   252
55298
53569ca831f4 allow merging of steps with subproofs
blanchet
parents: 55287
diff changeset
   253
              val (steps_before, _ :: steps_after) = chop i steps
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   254
              (* replace successors with their modified versions *)
55298
53569ca831f4 allow merging of steps with subproofs
blanchet
parents: 55287
diff changeset
   255
              val steps_after = update_steps steps_after succs'
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   256
            in
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   257
              decrement_step_count (); (* candidate successfully eliminated *)
55264
43473497fb65 centralize preplaying
blanchet
parents: 55263
diff changeset
   258
              map3 (fn time => set_preplay_outcomes_of_isar_step ctxt time preplay_data) times
55278
73372494fd80 more flexible compression, choosing whichever proof method works
blanchet
parents: 55272
diff changeset
   259
                succs' meths_outcomess;
55264
43473497fb65 centralize preplaying
blanchet
parents: 55263
diff changeset
   260
              map (replace_successor l labels) lfs;
55243
66709d41601e reset timing information after changes
blanchet
parents: 55223
diff changeset
   261
              (* removing the step would mess up the indices; replace with dummy step instead *)
55298
53569ca831f4 allow merging of steps with subproofs
blanchet
parents: 55287
diff changeset
   262
              steps_before @ dummy_isar_step :: steps_after
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   263
            end
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   264
            handle Bind => steps
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   265
                 | Match => steps
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   266
                 | Option.Option => steps
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   267
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   268
          fun compression_loop candidates steps =
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   269
            if not (compress_further ()) then
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   270
              steps
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   271
            else
55298
53569ca831f4 allow merging of steps with subproofs
blanchet
parents: 55287
diff changeset
   272
              (case pop_next_candidate candidates of
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   273
                (NONE, _) => steps (* no more candidates for elimination *)
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   274
              | (SOME (cand as (_, l, _)), candidates) =>
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   275
                let val successors = get_successors l in
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   276
                  if length successors > compress_degree then steps
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   277
                  else compression_loop candidates (try_eliminate cand successors steps)
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   278
                end)
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   279
        in
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   280
          compression_loop candidates steps
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   281
          |> remove (op =) dummy_isar_step
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   282
        end
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   283
55268
blanchet
parents: 55266
diff changeset
   284
      (* Proofs are compressed bottom-up, beginning with the innermost subproofs. On the innermost
blanchet
parents: 55266
diff changeset
   285
         proof level, the proof steps have no subproofs. In the best case, these steps can be merged
blanchet
parents: 55266
diff changeset
   286
         into just one step, resulting in a trivial subproof. Going one level up, trivial subproofs
blanchet
parents: 55266
diff changeset
   287
         can be eliminated. In the best case, this once again leads to a proof whose proof steps do
blanchet
parents: 55266
diff changeset
   288
         not have subproofs. Applying this approach recursively will result in a flat proof in the
blanchet
parents: 55266
diff changeset
   289
         best cast. *)
55263
blanchet
parents: 55260
diff changeset
   290
      fun compress_proof (proof as (Proof (fix, assms, steps))) =
blanchet
parents: 55260
diff changeset
   291
        if compress_further () then Proof (fix, assms, compress_steps steps) else proof
blanchet
parents: 55260
diff changeset
   292
      and compress_steps steps =
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   293
        (* bottom-up: compress innermost proofs first *)
55263
blanchet
parents: 55260
diff changeset
   294
        steps |> map (fn step => step |> compress_further () ? compress_sub_levels)
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   295
              |> compress_further () ? compress_top_level
55263
blanchet
parents: 55260
diff changeset
   296
      and compress_sub_levels (step as Let _) = step
55299
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
   297
        | 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: 52454
diff changeset
   298
          (* compress subproofs *)
55299
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
   299
          Prove (qs, xs, l, t, map compress_proof subproofs, facts, meths, comment)
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   300
          (* eliminate trivial subproofs *)
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   301
          |> compress_further () ? elim_subproofs
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   302
    in
55263
blanchet
parents: 55260
diff changeset
   303
      compress_proof proof
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   304
    end
50259
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
   305
54504
blanchet
parents: 54503
diff changeset
   306
end;