src/HOL/Tools/Sledgehammer/sledgehammer_compress.ML
author smolkas
Mon, 06 May 2013 11:03:08 +0200
changeset 51879 ee9562d31778
parent 51877 71052c42edf2
child 51930 52fd62618631
permissions -rw-r--r--
added preplay tracing
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
51130
76d68444cd59 renamed sledgehammer_shrink to sledgehammer_compress
smolkas
parents: 51128
diff changeset
     1
(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_compress.ML
50263
0b430064296a added comments to new source files
smolkas
parents: 50261
diff changeset
     2
    Author:     Jasmin Blanchette, TU Muenchen
0b430064296a added comments to new source files
smolkas
parents: 50261
diff changeset
     3
    Author:     Steffen Juilf Smolka, TU Muenchen
0b430064296a added comments to new source files
smolkas
parents: 50261
diff changeset
     4
51130
76d68444cd59 renamed sledgehammer_shrink to sledgehammer_compress
smolkas
parents: 51128
diff changeset
     5
Compression of reconstructed isar proofs.
50263
0b430064296a added comments to new source files
smolkas
parents: 50261
diff changeset
     6
*)
0b430064296a added comments to new source files
smolkas
parents: 50261
diff changeset
     7
51130
76d68444cd59 renamed sledgehammer_shrink to sledgehammer_compress
smolkas
parents: 51128
diff changeset
     8
signature SLEDGEHAMMER_COMPRESS =
50259
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
     9
sig
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
    10
  type isar_proof = Sledgehammer_Proof.isar_proof
50924
beb95bf66b21 changed type of preplay time; tuned preplaying
smolkas
parents: 50923
diff changeset
    11
  type preplay_time = Sledgehammer_Preplay.preplay_time
51741
blanchet
parents: 51260
diff changeset
    12
  val compress_and_preplay_proof :
50557
31313171deb5 thread no timeout properly
blanchet
parents: 50278
diff changeset
    13
    bool -> Proof.context -> string -> string -> bool -> Time.time option
51879
ee9562d31778 added preplay tracing
smolkas
parents: 51877
diff changeset
    14
    -> bool -> real -> isar_proof -> isar_proof * (bool * preplay_time)
50259
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    15
end
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    16
51130
76d68444cd59 renamed sledgehammer_shrink to sledgehammer_compress
smolkas
parents: 51128
diff changeset
    17
structure Sledgehammer_Compress : SLEDGEHAMMER_COMPRESS =
50259
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    18
struct
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    19
50265
9eafa567e061 made use of sledgehammer_util
smolkas
parents: 50264
diff changeset
    20
open Sledgehammer_Util
50264
a9ec48b98734 renamed sledgehammer_isar_reconstruct to sledgehammer_proof
smolkas
parents: 50263
diff changeset
    21
open Sledgehammer_Proof
50923
141d8f575f6f move preplaying to own structure
smolkas
parents: 50908
diff changeset
    22
open Sledgehammer_Preplay
50259
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    23
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    24
(* Parameters *)
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    25
val merge_timeout_slack = 1.2
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    26
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    27
(* Data structures, orders *)
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    28
val label_ord = prod_ord int_ord fast_string_ord o pairself swap
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    29
structure Label_Table = Table(
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    30
  type key = label
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    31
  val ord = label_ord)
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    32
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    33
(* clean vector interface *)
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    34
fun get i v = Vector.sub (v, i)
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    35
fun replace x i v = Vector.update (v, i, x)
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    36
fun update f i v = replace (get i v |> f) i v
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    37
fun v_fold_index f v s =
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    38
  Vector.foldl (fn (x, (i, s)) => (i+1, f (i, x) s)) (0, s) v |> snd
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    39
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    40
(* Queue interface to table *)
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    41
fun pop tab key =
51876
724c67f59929 added informative error messages
smolkas
parents: 51741
diff changeset
    42
  (let val v = hd (Inttab.lookup_list tab key) in
50259
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    43
    (v, Inttab.remove_list (op =) (key, v) tab)
51876
724c67f59929 added informative error messages
smolkas
parents: 51741
diff changeset
    44
  end) handle Empty => raise Fail "sledgehammer_compress: pop"
50259
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    45
fun pop_max tab = pop tab (the (Inttab.max_key tab))
51876
724c67f59929 added informative error messages
smolkas
parents: 51741
diff changeset
    46
  handle Option => raise Fail "sledgehammer_compress: pop_max"
50259
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    47
fun add_list tab xs = fold (Inttab.insert_list (op =)) xs tab
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    48
51130
76d68444cd59 renamed sledgehammer_shrink to sledgehammer_compress
smolkas
parents: 51128
diff changeset
    49
(* Main function for compresing proofs *)
51741
blanchet
parents: 51260
diff changeset
    50
fun compress_and_preplay_proof debug ctxt type_enc lam_trans preplay
51879
ee9562d31778 added preplay tracing
smolkas
parents: 51877
diff changeset
    51
      preplay_timeout preplay_trace isar_compress proof =
50259
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    52
  let
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
    53
    (* 60 seconds seems like a good interpreation of "no timeout" *)
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
    54
    val preplay_timeout = preplay_timeout |> the_default (seconds 60.0)
50271
2be84eaf7ebb deal with the case that metis does not time out, but fails instead
smolkas
parents: 50270
diff changeset
    55
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
    56
    (* handle metis preplay fail *)
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
    57
    local
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
    58
      open Unsynchronized
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
    59
      val metis_fail = ref false
50259
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    60
    in
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
    61
      fun handle_metis_fail try_metis () =
50901
wenzelm
parents: 50786
diff changeset
    62
        try_metis () handle exn =>
wenzelm
parents: 50786
diff changeset
    63
          (if Exn.is_interrupt exn orelse debug then reraise exn
50924
beb95bf66b21 changed type of preplay time; tuned preplaying
smolkas
parents: 50923
diff changeset
    64
           else metis_fail := true; some_preplay_time)
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
    65
      fun get_time lazy_time =
50924
beb95bf66b21 changed type of preplay time; tuned preplaying
smolkas
parents: 50923
diff changeset
    66
        if !metis_fail andalso not (Lazy.is_finished lazy_time)
beb95bf66b21 changed type of preplay time; tuned preplaying
smolkas
parents: 50923
diff changeset
    67
          then some_preplay_time
beb95bf66b21 changed type of preplay time; tuned preplaying
smolkas
parents: 50923
diff changeset
    68
          else Lazy.force lazy_time
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
    69
      val metis_fail = fn () => !metis_fail
50259
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    70
    end
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    71
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
    72
    (* compress top level steps - do not compress subproofs *)
51260
61bc5a3bef09 tuned agressiveness of isar compression
smolkas
parents: 51179
diff changeset
    73
    fun compress_top_level on_top_level ctxt n steps =
50259
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
    74
    let
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
    75
      (* proof step vector *)
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
    76
      val step_vect = steps |> map SOME |> Vector.fromList
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
    77
      val n_metis = add_metis_steps_top_level steps 0
51130
76d68444cd59 renamed sledgehammer_shrink to sledgehammer_compress
smolkas
parents: 51128
diff changeset
    78
      val target_n_metis = Real.fromInt n_metis / isar_compress |> Real.round
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
    79
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
    80
      (* table for mapping from (top-level-)label to step_vect position *)
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
    81
      fun update_table (i, Prove (_, l, _, _)) = Label_Table.update_new (l, i)
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
    82
        | update_table (i, Obtain (_, _, l, _, _)) = Label_Table.update_new (l, i)
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
    83
        | update_table _ = I
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
    84
      val label_index_table = fold_index update_table steps Label_Table.empty
50711
smolkas
parents: 50678
diff changeset
    85
      val lookup_indices = map_filter (Label_Table.lookup label_index_table)
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
    86
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
    87
      (* proof step references *)
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
    88
      fun refs step =
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
    89
        (case byline_of_step step of
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
    90
          NONE => []
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
    91
        | SOME (By_Metis (subproofs, (lfs, _))) =>
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
    92
            maps (steps_of_proof #> maps refs) subproofs @ lookup_indices lfs)
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
    93
      val refed_by_vect =
51260
61bc5a3bef09 tuned agressiveness of isar compression
smolkas
parents: 51179
diff changeset
    94
        Vector.tabulate (Vector.length step_vect, K [])
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
    95
        |> fold_index (fn (i, step) => fold (update (cons i)) (refs step)) steps
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
    96
        |> Vector.map rev (* after rev, indices are sorted in ascending order *)
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
    97
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
    98
      (* candidates for elimination, use table as priority queue (greedy
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
    99
         algorithm) *)
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   100
      fun add_if_cand step_vect (i, [j]) =
51876
724c67f59929 added informative error messages
smolkas
parents: 51741
diff changeset
   101
          ((case (the (get i step_vect), the (get j step_vect)) of
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   102
            (Prove (_, _, t, By_Metis _), Prove (_, _, _, By_Metis _)) =>
50780
4174abe2c5fd consider merging obtain steps
smolkas
parents: 50779
diff changeset
   103
              cons (Term.size_of_term t, i)
4174abe2c5fd consider merging obtain steps
smolkas
parents: 50779
diff changeset
   104
          | (Prove (_, _, t, By_Metis _), Obtain (_, _, _, _, By_Metis _)) =>
4174abe2c5fd consider merging obtain steps
smolkas
parents: 50779
diff changeset
   105
              cons (Term.size_of_term t, i)
51879
ee9562d31778 added preplay tracing
smolkas
parents: 51877
diff changeset
   106
          | _ => I)
51876
724c67f59929 added informative error messages
smolkas
parents: 51741
diff changeset
   107
            handle Option => raise Fail "sledgehammer_compress: add_if_cand")
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   108
        | add_if_cand _ _ = I
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   109
      val cand_tab =
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   110
        v_fold_index (add_if_cand step_vect) refed_by_vect []
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   111
        |> Inttab.make_list
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   112
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   113
      (* cache metis preplay times in lazy time vector *)
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   114
      val metis_time =
51178
06689dbfe072 simplified byline, isar_qualifier
smolkas
parents: 51147
diff changeset
   115
        Vector.map
50924
beb95bf66b21 changed type of preplay time; tuned preplaying
smolkas
parents: 50923
diff changeset
   116
          (if not preplay then K (zero_preplay_time) #> Lazy.value
50923
141d8f575f6f move preplaying to own structure
smolkas
parents: 50908
diff changeset
   117
           else
51178
06689dbfe072 simplified byline, isar_qualifier
smolkas
parents: 51147
diff changeset
   118
             the
51879
ee9562d31778 added preplay tracing
smolkas
parents: 51877
diff changeset
   119
             #> try_metis debug preplay_trace type_enc lam_trans ctxt
ee9562d31778 added preplay tracing
smolkas
parents: 51877
diff changeset
   120
              preplay_timeout
50923
141d8f575f6f move preplaying to own structure
smolkas
parents: 50908
diff changeset
   121
             #> handle_metis_fail
141d8f575f6f move preplaying to own structure
smolkas
parents: 50908
diff changeset
   122
             #> Lazy.lazy)
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   123
          step_vect
51876
724c67f59929 added informative error messages
smolkas
parents: 51741
diff changeset
   124
        handle Option => raise Fail "sledgehammer_compress: metis_time"
50923
141d8f575f6f move preplaying to own structure
smolkas
parents: 50908
diff changeset
   125
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   126
      fun sum_up_time lazy_time_vector =
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   127
        Vector.foldl
50924
beb95bf66b21 changed type of preplay time; tuned preplaying
smolkas
parents: 50923
diff changeset
   128
          (apfst get_time #> uncurry add_preplay_time)
beb95bf66b21 changed type of preplay time; tuned preplaying
smolkas
parents: 50923
diff changeset
   129
          zero_preplay_time lazy_time_vector
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   130
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   131
      (* Merging *)
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   132
      fun merge (Prove (_, lbl1, _, By_Metis (subproofs1, (lfs1, gfs1)))) step2 =
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   133
          let
51178
06689dbfe072 simplified byline, isar_qualifier
smolkas
parents: 51147
diff changeset
   134
            val (step_constructor, (subproofs2, (lfs2, gfs2))) =
50780
4174abe2c5fd consider merging obtain steps
smolkas
parents: 50779
diff changeset
   135
              (case step2 of
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   136
                Prove (qs2, lbl2, t, By_Metis x) =>
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   137
                  (fn by => Prove (qs2, lbl2, t, by), x)
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   138
              | Obtain (qs2, xs, lbl2, t, By_Metis x) =>
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   139
                  (fn by => Obtain (qs2, xs, lbl2, t, by), x)
51876
724c67f59929 added informative error messages
smolkas
parents: 51741
diff changeset
   140
              | _ => raise Fail "sledgehammer_compress: unmergeable Isar steps" )
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   141
            val lfs = remove (op =) lbl1 lfs2 |> union (op =) lfs1
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   142
            val gfs = union (op =) gfs1 gfs2
51178
06689dbfe072 simplified byline, isar_qualifier
smolkas
parents: 51147
diff changeset
   143
            val subproofs = subproofs1 @ subproofs2
06689dbfe072 simplified byline, isar_qualifier
smolkas
parents: 51147
diff changeset
   144
          in step_constructor (By_Metis (subproofs, (lfs, gfs))) end
51876
724c67f59929 added informative error messages
smolkas
parents: 51741
diff changeset
   145
        | merge _ _ = raise Fail "sledgehammer_compress: unmergeable Isar steps"
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   146
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   147
      fun try_merge metis_time (s1, i) (s2, j) =
50924
beb95bf66b21 changed type of preplay time; tuned preplaying
smolkas
parents: 50923
diff changeset
   148
        if not preplay then (merge s1 s2 |> SOME, metis_time)
beb95bf66b21 changed type of preplay time; tuned preplaying
smolkas
parents: 50923
diff changeset
   149
        else
beb95bf66b21 changed type of preplay time; tuned preplaying
smolkas
parents: 50923
diff changeset
   150
          (case get i metis_time |> Lazy.force of
beb95bf66b21 changed type of preplay time; tuned preplaying
smolkas
parents: 50923
diff changeset
   151
            (true, _) => (NONE, metis_time)
beb95bf66b21 changed type of preplay time; tuned preplaying
smolkas
parents: 50923
diff changeset
   152
          | (_, t1) =>
beb95bf66b21 changed type of preplay time; tuned preplaying
smolkas
parents: 50923
diff changeset
   153
            (case get j metis_time |> Lazy.force of
beb95bf66b21 changed type of preplay time; tuned preplaying
smolkas
parents: 50923
diff changeset
   154
              (true, _) => (NONE, metis_time)
beb95bf66b21 changed type of preplay time; tuned preplaying
smolkas
parents: 50923
diff changeset
   155
            | (_, t2) =>
beb95bf66b21 changed type of preplay time; tuned preplaying
smolkas
parents: 50923
diff changeset
   156
              let
beb95bf66b21 changed type of preplay time; tuned preplaying
smolkas
parents: 50923
diff changeset
   157
                val s12 = merge s1 s2
beb95bf66b21 changed type of preplay time; tuned preplaying
smolkas
parents: 50923
diff changeset
   158
                val timeout = time_mult merge_timeout_slack (Time.+(t1, t2))
beb95bf66b21 changed type of preplay time; tuned preplaying
smolkas
parents: 50923
diff changeset
   159
              in
51879
ee9562d31778 added preplay tracing
smolkas
parents: 51877
diff changeset
   160
                case try_metis_quietly debug preplay_trace type_enc
ee9562d31778 added preplay tracing
smolkas
parents: 51877
diff changeset
   161
                                                lam_trans ctxt timeout s12 () of
50924
beb95bf66b21 changed type of preplay time; tuned preplaying
smolkas
parents: 50923
diff changeset
   162
                  (true, _) => (NONE, metis_time)
beb95bf66b21 changed type of preplay time; tuned preplaying
smolkas
parents: 50923
diff changeset
   163
                | exact_time =>
beb95bf66b21 changed type of preplay time; tuned preplaying
smolkas
parents: 50923
diff changeset
   164
                  (SOME s12, metis_time
beb95bf66b21 changed type of preplay time; tuned preplaying
smolkas
parents: 50923
diff changeset
   165
                             |> replace (zero_preplay_time |> Lazy.value) i
beb95bf66b21 changed type of preplay time; tuned preplaying
smolkas
parents: 50923
diff changeset
   166
                             |> replace (Lazy.value exact_time) j)
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   167
50924
beb95bf66b21 changed type of preplay time; tuned preplaying
smolkas
parents: 50923
diff changeset
   168
              end))
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   169
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   170
      fun merge_steps metis_time step_vect refed_by cand_tab n' n_metis' =
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   171
        if Inttab.is_empty cand_tab
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   172
          orelse n_metis' <= target_n_metis
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   173
          orelse (on_top_level andalso n'<3)
51877
71052c42edf2 avoid dummy annotations; terminate preplay/compression if metis fails; fixed bug
smolkas
parents: 51876
diff changeset
   174
          orelse metis_fail()
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   175
        then
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   176
          (Vector.foldr
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   177
             (fn (NONE, steps) => steps | (SOME s, steps) => s :: steps)
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   178
             [] step_vect,
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   179
           sum_up_time metis_time)
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   180
        else
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   181
          let
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   182
            val (i, cand_tab) = pop_max cand_tab
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   183
            val j = get i refed_by |> the_single
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   184
            val s1 = get i step_vect |> the
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   185
            val s2 = get j step_vect |> the
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   186
          in
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   187
            case try_merge metis_time (s1, i) (s2, j) of
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   188
              (NONE, metis_time) =>
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   189
              merge_steps metis_time step_vect refed_by cand_tab n' n_metis'
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   190
            | (s, metis_time) =>
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   191
            let
51877
71052c42edf2 avoid dummy annotations; terminate preplay/compression if metis fails; fixed bug
smolkas
parents: 51876
diff changeset
   192
              val refs_s1 = refs s1
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   193
              val refed_by = refed_by |> fold
51877
71052c42edf2 avoid dummy annotations; terminate preplay/compression if metis fails; fixed bug
smolkas
parents: 51876
diff changeset
   194
                (update (Ord_List.remove int_ord i #> Ord_List.insert int_ord j))
71052c42edf2 avoid dummy annotations; terminate preplay/compression if metis fails; fixed bug
smolkas
parents: 51876
diff changeset
   195
                refs_s1
71052c42edf2 avoid dummy annotations; terminate preplay/compression if metis fails; fixed bug
smolkas
parents: 51876
diff changeset
   196
              val shared_refs = Ord_List.inter int_ord refs_s1 (refs s2)
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   197
              val new_candidates =
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   198
                fold (add_if_cand step_vect)
51877
71052c42edf2 avoid dummy annotations; terminate preplay/compression if metis fails; fixed bug
smolkas
parents: 51876
diff changeset
   199
                  (map (fn i => (i, get i refed_by)) shared_refs) []
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   200
              val cand_tab = add_list cand_tab new_candidates
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   201
              val step_vect = step_vect |> replace NONE i |> replace s j
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   202
            in
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   203
              merge_steps metis_time step_vect refed_by cand_tab (n' - 1)
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   204
                          (n_metis' - 1)
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   205
            end
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   206
          end
51876
724c67f59929 added informative error messages
smolkas
parents: 51741
diff changeset
   207
        handle Option => raise Fail "sledgehammer_compress: merge_steps"
724c67f59929 added informative error messages
smolkas
parents: 51741
diff changeset
   208
             | List.Empty => raise Fail "sledgehammer_compress: merge_steps"
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   209
    in
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   210
      merge_steps metis_time step_vect refed_by_vect cand_tab n n_metis
50259
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
   211
    end
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   212
51260
61bc5a3bef09 tuned agressiveness of isar compression
smolkas
parents: 51179
diff changeset
   213
    fun do_proof on_top_level ctxt (Proof (Fix fix, Assume assms, steps)) =
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   214
      let
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   215
        (* Enrich context with top-level facts *)
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   216
        val thy = Proof_Context.theory_of ctxt
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   217
        (* TODO: add Skolem variables to context? *)
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   218
        fun enrich_with_fact l t =
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   219
          Proof_Context.put_thms false
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   220
            (string_for_label l, SOME [Skip_Proof.make_thm thy t])
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   221
        fun enrich_with_step (Prove (_, l, t, _)) = enrich_with_fact l t
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   222
          | enrich_with_step (Obtain (_, _, l, t, _)) = enrich_with_fact l t
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   223
          | enrich_with_step _ = I
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   224
        val enrich_with_steps = fold enrich_with_step
51260
61bc5a3bef09 tuned agressiveness of isar compression
smolkas
parents: 51179
diff changeset
   225
        val enrich_with_assms = fold (uncurry enrich_with_fact)
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   226
        val rich_ctxt =
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   227
          ctxt |> enrich_with_assms assms |> enrich_with_steps steps
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   228
51260
61bc5a3bef09 tuned agressiveness of isar compression
smolkas
parents: 51179
diff changeset
   229
        val n = List.length fix + List.length assms + List.length steps
61bc5a3bef09 tuned agressiveness of isar compression
smolkas
parents: 51179
diff changeset
   230
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   231
        (* compress subproofs and top-levl steps *)
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   232
        val ((steps, top_level_time), lower_level_time) =
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   233
          steps |> do_subproofs rich_ctxt
51260
61bc5a3bef09 tuned agressiveness of isar compression
smolkas
parents: 51179
diff changeset
   234
                |>> compress_top_level on_top_level rich_ctxt n
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   235
      in
51260
61bc5a3bef09 tuned agressiveness of isar compression
smolkas
parents: 51179
diff changeset
   236
        (Proof (Fix fix, Assume assms, steps),
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   237
          add_preplay_time lower_level_time top_level_time)
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   238
      end
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   239
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   240
    and do_subproofs ctxt subproofs =
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   241
      let
51178
06689dbfe072 simplified byline, isar_qualifier
smolkas
parents: 51147
diff changeset
   242
        fun compress_each_and_collect_time compress subproofs =
06689dbfe072 simplified byline, isar_qualifier
smolkas
parents: 51147
diff changeset
   243
          let fun f_m proof time = compress proof ||> add_preplay_time time
06689dbfe072 simplified byline, isar_qualifier
smolkas
parents: 51147
diff changeset
   244
          in fold_map f_m subproofs zero_preplay_time end
06689dbfe072 simplified byline, isar_qualifier
smolkas
parents: 51147
diff changeset
   245
        val compress_subproofs =
51130
76d68444cd59 renamed sledgehammer_shrink to sledgehammer_compress
smolkas
parents: 51128
diff changeset
   246
          compress_each_and_collect_time (do_proof false ctxt)
51178
06689dbfe072 simplified byline, isar_qualifier
smolkas
parents: 51147
diff changeset
   247
        fun compress (Prove (qs, l, t, By_Metis(subproofs, facts))) =
06689dbfe072 simplified byline, isar_qualifier
smolkas
parents: 51147
diff changeset
   248
              let val (subproofs, time) = compress_subproofs subproofs
06689dbfe072 simplified byline, isar_qualifier
smolkas
parents: 51147
diff changeset
   249
              in (Prove (qs, l, t, By_Metis(subproofs, facts)), time) end
06689dbfe072 simplified byline, isar_qualifier
smolkas
parents: 51147
diff changeset
   250
          | compress atomic_step = (atomic_step, zero_preplay_time)
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   251
      in
51179
0d5f8812856f split isar_step into isar_step, fix, assms; made isar_proof explicit; register fixed variables in ctxt and auto_fix terms to avoid superfluous annotations
smolkas
parents: 51178
diff changeset
   252
        compress_each_and_collect_time compress subproofs
50672
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   253
      end
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   254
  in
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   255
    do_proof true ctxt proof
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   256
    |> apsnd (pair (metis_fail ()))
ab5b8b5c9cbe added "obtain" to Isar proof construction data structure
blanchet
parents: 50557
diff changeset
   257
  end
50259
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
   258
9c64a52ae499 put shrink in own structure
smolkas
parents:
diff changeset
   259
end