src/HOL/Tools/Sledgehammer/sledgehammer_isar_compress.ML
author fleury <Mathias.Fleury@mpi-inf.mpg.de>
Tue, 10 Nov 2015 17:49:54 +0100
changeset 61612 40859aa6d10c
parent 59058 a78612c67ec0
child 61666 f1b257607981
permissions -rw-r--r--
generalized so that is also works for veriT proofs
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
57054
blanchet
parents: 55452
diff changeset
    14
  val compress_isar_proof : Proof.context -> real -> Time.time ->
55329
3c2dbd2e221f more generous Isar proof compression -- try to remove failing steps
blanchet
parents: 55328
diff changeset
    15
    isar_preplay_data Unsynchronized.ref -> 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
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
    26
fun collect_successors steps lbls =
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
    27
  let
55329
3c2dbd2e221f more generous Isar proof compression -- try to remove failing steps
blanchet
parents: 55328
diff changeset
    28
    fun collect_steps _ (accum as ([], _)) = accum
55263
blanchet
parents: 55260
diff changeset
    29
      | collect_steps [] accum = accum
blanchet
parents: 55260
diff changeset
    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: 55328
diff changeset
    31
    and collect_step (step as Prove (_, _, l, _, subproofs, _, _, _)) x =
55263
blanchet
parents: 55260
diff changeset
    32
        (case collect_subproofs subproofs x of
55329
3c2dbd2e221f more generous Isar proof compression -- try to remove failing steps
blanchet
parents: 55328
diff changeset
    33
          (accum as ([], _)) => accum
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
    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: 55328
diff changeset
    35
        | collect_step _ accum = accum
3c2dbd2e221f more generous Isar proof compression -- try to remove failing steps
blanchet
parents: 55328
diff changeset
    36
    and collect_subproofs [] accum = accum
3c2dbd2e221f more generous Isar proof compression -- try to remove failing steps
blanchet
parents: 55328
diff changeset
    37
      | collect_subproofs (proof :: subproofs) accum =
3c2dbd2e221f more generous Isar proof compression -- try to remove failing steps
blanchet
parents: 55328
diff changeset
    38
        (case collect_steps (steps_of_isar_proof proof) accum of
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
    39
          accum as ([], _) => accum
55263
blanchet
parents: 55260
diff changeset
    40
        | accum => collect_subproofs subproofs accum)
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
    41
  in
55313
blanchet
parents: 55312
diff changeset
    42
    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
    43
  end
50259
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    44
55332
803a7400cc58 tuned code to avoid noncanonical (and risky) exception handling
blanchet
parents: 55331
diff changeset
    45
fun update_steps updates steps =
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
    46
  let
55263
blanchet
parents: 55260
diff changeset
    47
    fun update_steps [] updates = ([], updates)
blanchet
parents: 55260
diff changeset
    48
      | update_steps steps [] = (steps, [])
blanchet
parents: 55260
diff changeset
    49
      | update_steps (step :: steps) updates = update_step step (update_steps steps updates)
blanchet
parents: 55260
diff changeset
    50
    and update_step step (steps, []) = (step :: steps, [])
55299
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
    51
      | update_step (Prove (qs, xs, l, t, subproofs, facts, meths, comment))
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
    52
          (steps,
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
    53
           updates as Prove (qs', xs', l', t', subproofs', facts', meths', comment') :: updates') =
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
    54
        (if l = l' then
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
    55
           update_subproofs subproofs' updates'
57760
7f11f325c47d no need to 'obtain' variables not in formula
blanchet
parents: 57725
diff changeset
    56
           |>> (fn subproofs'' => Prove (qs', xs', l', t', subproofs'', facts', meths', comment'))
55299
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
    57
         else
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
    58
           update_subproofs subproofs updates
57760
7f11f325c47d no need to 'obtain' variables not in formula
blanchet
parents: 57725
diff changeset
    59
           |>> (fn subproofs' => Prove (qs, xs, l, t, subproofs', facts, meths, comment)))
55299
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
    60
        |>> (fn step => step :: steps)
55329
3c2dbd2e221f more generous Isar proof compression -- try to remove failing steps
blanchet
parents: 55328
diff changeset
    61
      | update_step step (steps, updates) = (step :: steps, updates)
55263
blanchet
parents: 55260
diff changeset
    62
    and update_subproofs [] updates = ([], updates)
blanchet
parents: 55260
diff changeset
    63
      | update_subproofs steps [] = (steps, [])
blanchet
parents: 55260
diff changeset
    64
      | update_subproofs (proof :: subproofs) updates =
blanchet
parents: 55260
diff changeset
    65
        update_proof proof (update_subproofs subproofs updates)
blanchet
parents: 55260
diff changeset
    66
    and update_proof proof (proofs, []) = (proof :: proofs, [])
57760
7f11f325c47d no need to 'obtain' variables not in formula
blanchet
parents: 57725
diff changeset
    67
      | update_proof (Proof (xs, assms, steps)) (proofs, updates) =
57699
a6cf197c1f1e also try 'metis' with 'full_types'
blanchet
parents: 57285
diff changeset
    68
        let val (steps', updates') = update_steps steps updates in
57760
7f11f325c47d no need to 'obtain' variables not in formula
blanchet
parents: 57725
diff changeset
    69
          (Proof (xs, assms, steps') :: proofs, updates')
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
    70
        end
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
    71
  in
55313
blanchet
parents: 55312
diff changeset
    72
    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
    73
  end
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff 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: 55280
diff 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: 55280
diff changeset
    76
  let
55328
0e17e92248ac tweaked handling of 'hopeless' methods
blanchet
parents: 55325
diff changeset
    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: 55280
diff 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: 55280
diff changeset
    79
        not (Lazy.is_finished outcome) orelse
55328
0e17e92248ac tweaked handling of 'hopeless' methods
blanchet
parents: 55325
diff changeset
    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: 55280
diff changeset
    81
      end
55328
0e17e92248ac tweaked handling of 'hopeless' methods
blanchet
parents: 55325
diff changeset
    82
0e17e92248ac tweaked handling of 'hopeless' methods
blanchet
parents: 55325
diff changeset
    83
    val (hopeful, hopeless) =
0e17e92248ac tweaked handling of 'hopeless' methods
blanchet
parents: 55325
diff changeset
    84
      meths2 @ subtract (op =) meths2 meths1
0e17e92248ac tweaked handling of 'hopeless' methods
blanchet
parents: 55325
diff changeset
    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: 55280
diff changeset
    86
  in
55328
0e17e92248ac tweaked handling of 'hopeless' methods
blanchet
parents: 55325
diff changeset
    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: 55280
diff 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: 55280
diff changeset
    89
57760
7f11f325c47d no need to 'obtain' variables not in formula
blanchet
parents: 57725
diff changeset
    90
fun merge_steps preplay_data (Prove ([], xs1, l1, _, subproofs1, (lfs1, gfs1), meths1, comment1))
7f11f325c47d no need to 'obtain' variables not in formula
blanchet
parents: 57725
diff changeset
    91
      (Prove (qs2, xs2, l2, t, subproofs2, (lfs2, gfs2), meths2, comment2)) =
55332
803a7400cc58 tuned code to avoid noncanonical (and risky) exception handling
blanchet
parents: 55331
diff changeset
    92
  let
803a7400cc58 tuned code to avoid noncanonical (and risky) exception handling
blanchet
parents: 55331
diff 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: 55331
diff changeset
    94
    val lfs = union (op =) lfs1 (remove (op =) l1 lfs2)
803a7400cc58 tuned code to avoid noncanonical (and risky) exception handling
blanchet
parents: 55331
diff changeset
    95
    val gfs = union (op =) gfs1 gfs2
803a7400cc58 tuned code to avoid noncanonical (and risky) exception handling
blanchet
parents: 55331
diff changeset
    96
  in
57760
7f11f325c47d no need to 'obtain' variables not in formula
blanchet
parents: 57725
diff changeset
    97
    (Prove (qs2, inter (op =) (Term.add_frees t []) (xs1 @ xs2), l2, t,
57776
1111a9a328fe rationalized sorting of facts -- so that preplaying (almost always) coincides with the real thing, preventing odd failures
blanchet
parents: 57768
diff changeset
    98
       subproofs1 @ subproofs2, sort_facts (lfs, gfs), meths, comment1 ^ comment2),
1111a9a328fe rationalized sorting of facts -- so that preplaying (almost always) coincides with the real thing, preventing odd failures
blanchet
parents: 57768
diff changeset
    99
     hopeless)
55332
803a7400cc58 tuned code to avoid noncanonical (and risky) exception handling
blanchet
parents: 55331
diff changeset
   100
  end
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   101
57768
a63f14f1214c fine-tuned Isar reconstruction, esp. boolean simplifications
blanchet
parents: 57766
diff changeset
   102
val merge_slack_time = seconds 0.01
55329
3c2dbd2e221f more generous Isar proof compression -- try to remove failing steps
blanchet
parents: 55328
diff changeset
   103
val merge_slack_factor = 1.5
3c2dbd2e221f more generous Isar proof compression -- try to remove failing steps
blanchet
parents: 55328
diff changeset
   104
3c2dbd2e221f more generous Isar proof compression -- try to remove failing steps
blanchet
parents: 55328
diff changeset
   105
fun adjust_merge_timeout max time =
3c2dbd2e221f more generous Isar proof compression -- try to remove failing steps
blanchet
parents: 55328
diff changeset
   106
  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: 55328
diff changeset
   107
    if Time.< (max, timeout) then max else timeout
3c2dbd2e221f more generous Isar proof compression -- try to remove failing steps
blanchet
parents: 55328
diff changeset
   108
  end
3c2dbd2e221f more generous Isar proof compression -- try to remove failing steps
blanchet
parents: 55328
diff changeset
   109
53763
70d370743dc6 hardcoded obscure option
blanchet
parents: 53762
diff changeset
   110
val compress_degree = 2
55271
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. *)
57245
f6bf6d5341ee renamed Sledgehammer options
blanchet
parents: 57162
diff changeset
   113
fun compress_isar_proof ctxt compress preplay_timeout preplay_data proof =
f6bf6d5341ee renamed Sledgehammer options
blanchet
parents: 57162
diff changeset
   114
  if compress <= 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
57768
a63f14f1214c fine-tuned Isar reconstruction, esp. boolean simplifications
blanchet
parents: 57766
diff changeset
   121
          val target_number_of_steps = Real.ceil (Real.fromInt number_of_steps / compress)
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
55333
fa079fd40db8 more exception cleanup + more liberal compressions of steps that timed out
blanchet
parents: 55332
diff changeset
   151
      fun reference_time l =
fa079fd40db8 more exception cleanup + more liberal compressions of steps that timed out
blanchet
parents: 55332
diff changeset
   152
        (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: 55332
diff changeset
   153
          Played time => time
fa079fd40db8 more exception cleanup + more liberal compressions of steps that timed out
blanchet
parents: 55332
diff changeset
   154
        | _ => preplay_timeout)
fa079fd40db8 more exception cleanup + more liberal compressions of steps that timed out
blanchet
parents: 55332
diff changeset
   155
55243
66709d41601e reset timing information after changes
blanchet
parents: 55223
diff changeset
   156
      (* elimination of trivial, one-step subproofs *)
57760
7f11f325c47d no need to 'obtain' variables not in formula
blanchet
parents: 57725
diff changeset
   157
      fun elim_one_subproof time (step as Prove (qs, xs, l, t, _, (lfs, gfs), meths, comment)) subs
55309
455a7f9924df don't lose additional outcomes
blanchet
parents: 55307
diff changeset
   158
          nontriv_subs =
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   159
        if null subs orelse not (compress_further ()) then
57760
7f11f325c47d no need to 'obtain' variables not in formula
blanchet
parents: 57725
diff changeset
   160
          Prove (qs, xs, l, t, List.revAppend (nontriv_subs, subs), (lfs, gfs), meths, comment)
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   161
        else
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   162
          (case subs of
55333
fa079fd40db8 more exception cleanup + more liberal compressions of steps that timed out
blanchet
parents: 55332
diff changeset
   163
            (sub as Proof (_, assms, [Prove (_, _, l', _, [], (lfs', gfs'), meths', _)])) :: subs =>
fa079fd40db8 more exception cleanup + more liberal compressions of steps that timed out
blanchet
parents: 55332
diff changeset
   164
            let
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   165
              (* merge steps *)
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   166
              val subs'' = subs @ nontriv_subs
55223
3c593bad6b31 generalized preplaying infrastructure to store various results for various methods
blanchet
parents: 55221
diff changeset
   167
              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
   168
              val gfs'' = union (op =) gfs' gfs
55328
0e17e92248ac tweaked handling of 'hopeless' methods
blanchet
parents: 55325
diff changeset
   169
              val (meths'' as _ :: _, hopeless) =
0e17e92248ac tweaked handling of 'hopeless' methods
blanchet
parents: 55325
diff changeset
   170
                merge_methods (!preplay_data) (l', meths') (l, meths)
57760
7f11f325c47d no need to 'obtain' variables not in formula
blanchet
parents: 57725
diff changeset
   171
              val step'' = Prove (qs, xs, 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
   172
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   173
              (* 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: 55332
diff changeset
   174
              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: 52454
diff changeset
   175
            in
57054
blanchet
parents: 55452
diff changeset
   176
              (case preplay_isar_step ctxt timeout hopeless step'' of
55333
fa079fd40db8 more exception cleanup + more liberal compressions of steps that timed out
blanchet
parents: 55332
diff changeset
   177
                meths_outcomes as (_, Played time'') :: _ =>
57725
a297879fe5b8 cascading timeout in parallel evaluation, to rapidly find optimum
blanchet
parents: 57699
diff changeset
   178
                (* "l'" successfully eliminated *)
55333
fa079fd40db8 more exception cleanup + more liberal compressions of steps that timed out
blanchet
parents: 55332
diff changeset
   179
                (decrement_step_count ();
57054
blanchet
parents: 55452
diff changeset
   180
                 set_preplay_outcomes_of_isar_step ctxt time'' preplay_data step'' meths_outcomes;
55333
fa079fd40db8 more exception cleanup + more liberal compressions of steps that timed out
blanchet
parents: 55332
diff changeset
   181
                 map (replace_successor l' [l]) lfs';
fa079fd40db8 more exception cleanup + more liberal compressions of steps that timed out
blanchet
parents: 55332
diff changeset
   182
                 elim_one_subproof time'' step'' subs nontriv_subs)
fa079fd40db8 more exception cleanup + more liberal compressions of steps that timed out
blanchet
parents: 55332
diff 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: 52454
diff changeset
   184
            end
55333
fa079fd40db8 more exception cleanup + more liberal compressions of steps that timed out
blanchet
parents: 55332
diff 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: 52454
diff changeset
   186
55333
fa079fd40db8 more exception cleanup + more liberal compressions of steps that timed out
blanchet
parents: 55332
diff changeset
   187
      fun elim_subproofs (step as Prove (_, _, l, _, subproofs, _, _, _)) =
fa079fd40db8 more exception cleanup + more liberal compressions of steps that timed out
blanchet
parents: 55332
diff 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: 55332
diff changeset
   189
            elim_one_subproof (reference_time l) step subproofs []
fa079fd40db8 more exception cleanup + more liberal compressions of steps that timed out
blanchet
parents: 55332
diff changeset
   190
          else
fa079fd40db8 more exception cleanup + more liberal compressions of steps that timed out
blanchet
parents: 55332
diff changeset
   191
            step
55299
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
   192
        | elim_subproofs step = step
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   193
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   194
      fun compress_top_level steps =
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   195
        let
57764
cac309e2b1f7 try to get rid of skolems first
blanchet
parents: 57760
diff changeset
   196
          val cand_key = apfst (length o get_successors)
cac309e2b1f7 try to get rid of skolems first
blanchet
parents: 57760
diff changeset
   197
          val cand_ord =
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58634
diff changeset
   198
            prod_ord int_ord (prod_ord (int_ord o swap) (int_ord o swap)) o apply2 cand_key
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   199
55298
53569ca831f4 allow merging of steps with subproofs
blanchet
parents: 55287
diff changeset
   200
          fun pop_next_candidate [] = (NONE, [])
53569ca831f4 allow merging of steps with subproofs
blanchet
parents: 55287
diff changeset
   201
            | pop_next_candidate (cands as (cand :: cands')) =
55331
c7561e87cba7 got rid of indices
blanchet
parents: 55330
diff changeset
   202
              fold (fn x => fn y => if cand_ord (x, y) = LESS then x else y) cands' cand
55325
462ffd3b7065 tuned code
blanchet
parents: 55323
diff changeset
   203
              |> (fn best => (SOME best, remove (op =) best cands))
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   204
55332
803a7400cc58 tuned code to avoid noncanonical (and risky) exception handling
blanchet
parents: 55331
diff changeset
   205
          fun try_eliminate i l labels steps =
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   206
            let
57776
1111a9a328fe rationalized sorting of facts -- so that preplaying (almost always) coincides with the real thing, preventing odd failures
blanchet
parents: 57768
diff changeset
   207
              val (steps_before, (cand as Prove (_, _, _, _, _, (lfs, _), _, _)) :: steps_after) =
55332
803a7400cc58 tuned code to avoid noncanonical (and risky) exception handling
blanchet
parents: 55331
diff changeset
   208
                chop i steps
55330
547d23e2abf7 corrected wrong 'meth :: _' pattern maching + tuned code
blanchet
parents: 55329
diff changeset
   209
              val succs = collect_successors steps_after labels
55332
803a7400cc58 tuned code to avoid noncanonical (and risky) exception handling
blanchet
parents: 55331
diff changeset
   210
              val (succs', hopelesses) = split_list (map (merge_steps (!preplay_data) cand) succs)
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   211
            in
55332
803a7400cc58 tuned code to avoid noncanonical (and risky) exception handling
blanchet
parents: 55331
diff changeset
   212
              (case try (map ((fn Played time => time) o
803a7400cc58 tuned code to avoid noncanonical (and risky) exception handling
blanchet
parents: 55331
diff changeset
   213
                  forced_intermediate_preplay_outcome_of_isar_step (!preplay_data))) labels of
803a7400cc58 tuned code to avoid noncanonical (and risky) exception handling
blanchet
parents: 55331
diff changeset
   214
                NONE => steps
803a7400cc58 tuned code to avoid noncanonical (and risky) exception handling
blanchet
parents: 55331
diff changeset
   215
              | SOME times0 =>
803a7400cc58 tuned code to avoid noncanonical (and risky) exception handling
blanchet
parents: 55331
diff changeset
   216
                let
61612
40859aa6d10c generalized so that is also works for veriT proofs
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 59058
diff changeset
   217
                  val n = length labels
40859aa6d10c generalized so that is also works for veriT proofs
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 59058
diff changeset
   218
                  val total_time = Library.foldl (op Time.+) (reference_time l, times0)
40859aa6d10c generalized so that is also works for veriT proofs
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 59058
diff changeset
   219
                  val timeout = adjust_merge_timeout preplay_timeout
40859aa6d10c generalized so that is also works for veriT proofs
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 59058
diff changeset
   220
                    (Time.fromReal (Time.toReal total_time / Real.fromInt n))
40859aa6d10c generalized so that is also works for veriT proofs
fleury <Mathias.Fleury@mpi-inf.mpg.de>
parents: 59058
diff changeset
   221
                  val meths_outcomess = @{map 2} (preplay_isar_step ctxt timeout) hopelesses succs'
55332
803a7400cc58 tuned code to avoid noncanonical (and risky) exception handling
blanchet
parents: 55331
diff changeset
   222
                in
803a7400cc58 tuned code to avoid noncanonical (and risky) exception handling
blanchet
parents: 55331
diff changeset
   223
                  (case try (map (fn (_, Played time) :: _ => time)) meths_outcomess of
803a7400cc58 tuned code to avoid noncanonical (and risky) exception handling
blanchet
parents: 55331
diff changeset
   224
                    NONE => steps
803a7400cc58 tuned code to avoid noncanonical (and risky) exception handling
blanchet
parents: 55331
diff changeset
   225
                  | SOME times =>
57725
a297879fe5b8 cascading timeout in parallel evaluation, to rapidly find optimum
blanchet
parents: 57699
diff changeset
   226
                    (* "l" successfully eliminated *)
55332
803a7400cc58 tuned code to avoid noncanonical (and risky) exception handling
blanchet
parents: 55331
diff changeset
   227
                    (decrement_step_count ();
58634
9f10d82e8188 added parameterized ML antiquotations @{map N}, @{fold N}, @{fold_map N}, @{split_list N};
wenzelm
parents: 57776
diff changeset
   228
                     @{map 3} (fn time => set_preplay_outcomes_of_isar_step ctxt time preplay_data)
55332
803a7400cc58 tuned code to avoid noncanonical (and risky) exception handling
blanchet
parents: 55331
diff changeset
   229
                       times succs' meths_outcomess;
803a7400cc58 tuned code to avoid noncanonical (and risky) exception handling
blanchet
parents: 55331
diff changeset
   230
                     map (replace_successor l labels) lfs;
803a7400cc58 tuned code to avoid noncanonical (and risky) exception handling
blanchet
parents: 55331
diff changeset
   231
                     steps_before @ update_steps succs' steps_after))
803a7400cc58 tuned code to avoid noncanonical (and risky) exception handling
blanchet
parents: 55331
diff changeset
   232
                end)
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   233
            end
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   234
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   235
          fun compression_loop candidates steps =
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   236
            if not (compress_further ()) then
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   237
              steps
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   238
            else
55298
53569ca831f4 allow merging of steps with subproofs
blanchet
parents: 55287
diff changeset
   239
              (case pop_next_candidate candidates of
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   240
                (NONE, _) => steps (* no more candidates for elimination *)
57766
77b48e7c0d89 careful when compressing 'obtains'
blanchet
parents: 57764
diff changeset
   241
              | (SOME (l, (num_xs, _)), candidates') =>
55332
803a7400cc58 tuned code to avoid noncanonical (and risky) exception handling
blanchet
parents: 55331
diff 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: 55331
diff changeset
   243
                  ~1 => steps
803a7400cc58 tuned code to avoid noncanonical (and risky) exception handling
blanchet
parents: 55331
diff changeset
   244
                | i =>
57766
77b48e7c0d89 careful when compressing 'obtains'
blanchet
parents: 57764
diff changeset
   245
                  let
77b48e7c0d89 careful when compressing 'obtains'
blanchet
parents: 57764
diff changeset
   246
                    val successors = get_successors l
77b48e7c0d89 careful when compressing 'obtains'
blanchet
parents: 57764
diff changeset
   247
                    val num_successors = length successors
77b48e7c0d89 careful when compressing 'obtains'
blanchet
parents: 57764
diff changeset
   248
                  in
77b48e7c0d89 careful when compressing 'obtains'
blanchet
parents: 57764
diff changeset
   249
                    (* Careful with "obtain", so we don't "obtain" twice the same variable after a
77b48e7c0d89 careful when compressing 'obtains'
blanchet
parents: 57764
diff changeset
   250
                       merge. *)
77b48e7c0d89 careful when compressing 'obtains'
blanchet
parents: 57764
diff changeset
   251
                    if num_successors > (if num_xs > 0 then 1 else compress_degree) then
57162
5ed907407041 avoid division by 0
blanchet
parents: 57054
diff changeset
   252
                      steps
5ed907407041 avoid division by 0
blanchet
parents: 57054
diff changeset
   253
                    else
5ed907407041 avoid division by 0
blanchet
parents: 57054
diff changeset
   254
                      steps
5ed907407041 avoid division by 0
blanchet
parents: 57054
diff changeset
   255
                      |> not (null successors) ? try_eliminate i l successors
5ed907407041 avoid division by 0
blanchet
parents: 57054
diff changeset
   256
                      |> compression_loop candidates'
55332
803a7400cc58 tuned code to avoid noncanonical (and risky) exception handling
blanchet
parents: 55331
diff changeset
   257
                  end))
55331
c7561e87cba7 got rid of indices
blanchet
parents: 55330
diff changeset
   258
57764
cac309e2b1f7 try to get rid of skolems first
blanchet
parents: 57760
diff changeset
   259
          fun add_cand (Prove (_, xs, l, t, _, _, _, _)) = cons (l, (length xs, size_of_term t))
55331
c7561e87cba7 got rid of indices
blanchet
parents: 55330
diff changeset
   260
            | add_cand _ = I
c7561e87cba7 got rid of indices
blanchet
parents: 55330
diff changeset
   261
c7561e87cba7 got rid of indices
blanchet
parents: 55330
diff changeset
   262
          (* the very last step is not a candidate *)
c7561e87cba7 got rid of indices
blanchet
parents: 55330
diff changeset
   263
          val candidates = fold add_cand (fst (split_last steps)) []
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   264
        in
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   265
          compression_loop candidates steps
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   266
        end
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   267
55268
blanchet
parents: 55266
diff changeset
   268
      (* Proofs are compressed bottom-up, beginning with the innermost subproofs. On the innermost
blanchet
parents: 55266
diff changeset
   269
         proof level, the proof steps have no subproofs. In the best case, these steps can be merged
blanchet
parents: 55266
diff changeset
   270
         into just one step, resulting in a trivial subproof. Going one level up, trivial subproofs
blanchet
parents: 55266
diff changeset
   271
         can be eliminated. In the best case, this once again leads to a proof whose proof steps do
blanchet
parents: 55266
diff changeset
   272
         not have subproofs. Applying this approach recursively will result in a flat proof in the
blanchet
parents: 55266
diff changeset
   273
         best cast. *)
57760
7f11f325c47d no need to 'obtain' variables not in formula
blanchet
parents: 57725
diff changeset
   274
      fun compress_proof (proof as (Proof (xs, assms, steps))) =
7f11f325c47d no need to 'obtain' variables not in formula
blanchet
parents: 57725
diff changeset
   275
        if compress_further () then Proof (xs, assms, compress_steps steps) else proof
55263
blanchet
parents: 55260
diff changeset
   276
      and compress_steps steps =
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   277
        (* bottom-up: compress innermost proofs first *)
57768
a63f14f1214c fine-tuned Isar reconstruction, esp. boolean simplifications
blanchet
parents: 57766
diff changeset
   278
        steps
a63f14f1214c fine-tuned Isar reconstruction, esp. boolean simplifications
blanchet
parents: 57766
diff changeset
   279
        |> map (fn step => step |> compress_further () ? compress_sub_levels)
a63f14f1214c fine-tuned Isar reconstruction, esp. boolean simplifications
blanchet
parents: 57766
diff changeset
   280
        |> compress_further () ? compress_top_level
55329
3c2dbd2e221f more generous Isar proof compression -- try to remove failing steps
blanchet
parents: 55328
diff changeset
   281
      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: 52454
diff changeset
   282
          (* compress subproofs *)
55299
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55298
diff changeset
   283
          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
   284
          (* eliminate trivial subproofs *)
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   285
          |> compress_further () ? elim_subproofs
55329
3c2dbd2e221f more generous Isar proof compression -- try to remove failing steps
blanchet
parents: 55328
diff changeset
   286
        | compress_sub_levels step = step
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   287
    in
55263
blanchet
parents: 55260
diff changeset
   288
      compress_proof proof
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
   289
    end
50259
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
   290
54504
blanchet
parents: 54503
diff changeset
   291
end;