src/HOL/Tools/Sledgehammer/sledgehammer_isar_preplay.ML
author desharna
Thu, 29 Oct 2020 16:07:41 +0100
changeset 72518 4be6ae020fc4
parent 69593 3dda49e08b9d
child 72582 b69a3a7655f2
permissions -rw-r--r--
Added smt (verit) to Sledgehammer's proof preplay. Tuned preplay multithreading.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
55202
824c48a539c9 renamed many Sledgehammer ML files to clarify structure
blanchet
parents: 55194
diff changeset
     1
(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_isar_preplay.ML
54712
cbebe2cf77f1 more work on Z3 Isar proofs
blanchet
parents: 54700
diff changeset
     2
    Author:     Steffen Juilf Smolka, TU Muenchen
50923
141d8f575f6f move preplaying to own structure
smolkas
parents:
diff changeset
     3
    Author:     Jasmin Blanchette, TU Muenchen
141d8f575f6f move preplaying to own structure
smolkas
parents:
diff changeset
     4
54763
blanchet
parents: 54761
diff changeset
     5
Preplaying of Isar proofs.
50923
141d8f575f6f move preplaying to own structure
smolkas
parents:
diff changeset
     6
*)
141d8f575f6f move preplaying to own structure
smolkas
parents:
diff changeset
     7
55202
824c48a539c9 renamed many Sledgehammer ML files to clarify structure
blanchet
parents: 55194
diff changeset
     8
signature SLEDGEHAMMER_ISAR_PREPLAY =
50923
141d8f575f6f move preplaying to own structure
smolkas
parents:
diff changeset
     9
sig
55287
ffa306239316 renamed ML file
blanchet
parents: 55285
diff changeset
    10
  type play_outcome = Sledgehammer_Proof_Methods.play_outcome
58426
blanchet
parents: 58079
diff changeset
    11
  type proof_method = Sledgehammer_Proof_Methods.proof_method
55212
blanchet
parents: 55202
diff changeset
    12
  type isar_step = Sledgehammer_Isar_Proof.isar_step
55202
824c48a539c9 renamed many Sledgehammer ML files to clarify structure
blanchet
parents: 55194
diff changeset
    13
  type isar_proof = Sledgehammer_Isar_Proof.isar_proof
824c48a539c9 renamed many Sledgehammer ML files to clarify structure
blanchet
parents: 55194
diff changeset
    14
  type label = Sledgehammer_Isar_Proof.label
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
    15
55212
blanchet
parents: 55202
diff changeset
    16
  val trace : bool Config.T
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
    17
55260
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
    18
  type isar_preplay_data
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
    19
58079
df0d6ce8fb66 made trace more informative when minimization is enabled
blanchet
parents: 57775
diff changeset
    20
  val peek_at_outcome : play_outcome Lazy.lazy -> play_outcome
55256
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
    21
  val enrich_context_with_local_facts : isar_proof -> Proof.context -> Proof.context
62219
dbac573b27e7 preplaying of 'smt' and 'metis' more in sync with actual method
blanchet
parents: 61268
diff changeset
    22
  val preplay_isar_step_for_method : Proof.context -> thm list -> Time.time -> proof_method ->
dbac573b27e7 preplaying of 'smt' and 'metis' more in sync with actual method
blanchet
parents: 61268
diff changeset
    23
    isar_step -> play_outcome
dbac573b27e7 preplaying of 'smt' and 'metis' more in sync with actual method
blanchet
parents: 61268
diff changeset
    24
  val preplay_isar_step : Proof.context -> thm list -> Time.time -> proof_method list ->
dbac573b27e7 preplaying of 'smt' and 'metis' more in sync with actual method
blanchet
parents: 61268
diff changeset
    25
    isar_step -> (proof_method * play_outcome) list
57054
blanchet
parents: 56093
diff changeset
    26
  val set_preplay_outcomes_of_isar_step : Proof.context -> Time.time ->
55278
73372494fd80 more flexible compression, choosing whichever proof method works
blanchet
parents: 55275
diff changeset
    27
    isar_preplay_data Unsynchronized.ref -> isar_step -> (proof_method * play_outcome) list -> unit
55272
blanchet
parents: 55269
diff changeset
    28
  val forced_intermediate_preplay_outcome_of_isar_step : isar_preplay_data -> label -> play_outcome
55266
d9d31354834e centralize more preplaying
blanchet
parents: 55264
diff changeset
    29
  val preplay_outcome_of_isar_step_for_method : isar_preplay_data -> label -> proof_method ->
55260
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
    30
    play_outcome Lazy.lazy
55266
d9d31354834e centralize more preplaying
blanchet
parents: 55264
diff changeset
    31
  val fastest_method_of_isar_step : isar_preplay_data -> label -> proof_method
55260
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
    32
  val preplay_outcome_of_isar_proof : isar_preplay_data -> isar_proof -> play_outcome
54504
blanchet
parents: 53761
diff changeset
    33
end;
50923
141d8f575f6f move preplaying to own structure
smolkas
parents:
diff changeset
    34
55202
824c48a539c9 renamed many Sledgehammer ML files to clarify structure
blanchet
parents: 55194
diff changeset
    35
structure Sledgehammer_Isar_Preplay : SLEDGEHAMMER_ISAR_PREPLAY =
50923
141d8f575f6f move preplaying to own structure
smolkas
parents:
diff changeset
    36
struct
141d8f575f6f move preplaying to own structure
smolkas
parents:
diff changeset
    37
55257
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55256
diff changeset
    38
open ATP_Proof_Reconstruct
50923
141d8f575f6f move preplaying to own structure
smolkas
parents:
diff changeset
    39
open Sledgehammer_Util
55287
ffa306239316 renamed ML file
blanchet
parents: 55285
diff changeset
    40
open Sledgehammer_Proof_Methods
55202
824c48a539c9 renamed many Sledgehammer ML files to clarify structure
blanchet
parents: 55194
diff changeset
    41
open Sledgehammer_Isar_Proof
50923
141d8f575f6f move preplaying to own structure
smolkas
parents:
diff changeset
    42
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 62826
diff changeset
    43
val trace = Attrib.setup_config_bool \<^binding>\<open>sledgehammer_preplay_trace\<close> (K false)
50924
beb95bf66b21 changed type of preplay time; tuned preplaying
smolkas
parents: 50923
diff changeset
    44
58079
df0d6ce8fb66 made trace more informative when minimization is enabled
blanchet
parents: 57775
diff changeset
    45
fun peek_at_outcome outcome =
df0d6ce8fb66 made trace more informative when minimization is enabled
blanchet
parents: 57775
diff changeset
    46
  if Lazy.is_finished outcome then Lazy.force outcome else Play_Timed_Out Time.zeroTime
df0d6ce8fb66 made trace more informative when minimization is enabled
blanchet
parents: 57775
diff changeset
    47
57775
40eea08c0cc5 slightly earlier exit from preplaying
blanchet
parents: 57734
diff changeset
    48
fun par_list_map_filter_with_timeout _ _ _ _ [] = []
40eea08c0cc5 slightly earlier exit from preplaying
blanchet
parents: 57734
diff changeset
    49
  | par_list_map_filter_with_timeout get_time min_timeout timeout0 f xs =
57734
18bb3e1ff6f6 rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents: 57725
diff changeset
    50
    let
18bb3e1ff6f6 rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents: 57725
diff changeset
    51
      val next_timeout = Unsynchronized.ref timeout0
57725
a297879fe5b8 cascading timeout in parallel evaluation, to rapidly find optimum
blanchet
parents: 57054
diff changeset
    52
57734
18bb3e1ff6f6 rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents: 57725
diff changeset
    53
      fun apply_f x =
18bb3e1ff6f6 rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents: 57725
diff changeset
    54
        let val timeout = !next_timeout in
62826
eb94e570c1a4 prefer infix operations;
wenzelm
parents: 62519
diff changeset
    55
          if timeout <= min_timeout then
57734
18bb3e1ff6f6 rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents: 57725
diff changeset
    56
            NONE
18bb3e1ff6f6 rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents: 57725
diff changeset
    57
          else
18bb3e1ff6f6 rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents: 57725
diff changeset
    58
            let val y = f timeout x in
18bb3e1ff6f6 rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents: 57725
diff changeset
    59
              (case get_time y of
72518
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 69593
diff changeset
    60
                SOME time => next_timeout := time_min (time, !next_timeout)
57734
18bb3e1ff6f6 rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents: 57725
diff changeset
    61
              | _ => ());
18bb3e1ff6f6 rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents: 57725
diff changeset
    62
              SOME y
18bb3e1ff6f6 rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents: 57725
diff changeset
    63
            end
18bb3e1ff6f6 rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents: 57725
diff changeset
    64
        end
18bb3e1ff6f6 rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents: 57725
diff changeset
    65
    in
72518
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 69593
diff changeset
    66
      chop_groups (Multithreading.max_threads ()) xs
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 69593
diff changeset
    67
      |> map (map_filter I o Par_List.map apply_f)
4be6ae020fc4 Added smt (verit) to Sledgehammer's proof preplay.
desharna
parents: 69593
diff changeset
    68
      |> flat
57734
18bb3e1ff6f6 rationalized preplaying by eliminating (now superfluous) laziness
blanchet
parents: 57725
diff changeset
    69
    end
57725
a297879fe5b8 cascading timeout in parallel evaluation, to rapidly find optimum
blanchet
parents: 57054
diff changeset
    70
55256
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
    71
fun enrich_context_with_local_facts proof ctxt =
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
    72
  let
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
    73
    val thy = Proof_Context.theory_of ctxt
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
    74
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
    75
    fun enrich_with_fact l t =
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
    76
      Proof_Context.put_thms false (string_of_label l, SOME [Skip_Proof.make_thm thy t])
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
    77
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
    78
    val enrich_with_assms = fold (uncurry enrich_with_fact)
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
    79
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
    80
    fun enrich_with_proof (Proof (_, assms, isar_steps)) =
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
    81
      enrich_with_assms assms #> fold enrich_with_step isar_steps
55299
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55295
diff changeset
    82
    and enrich_with_step (Prove (_, _, l, t, subproofs, _, _, _)) =
55256
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
    83
        enrich_with_fact l t #> fold enrich_with_proof subproofs
55299
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55295
diff changeset
    84
      | enrich_with_step _ = I
55256
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
    85
  in
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
    86
    enrich_with_proof proof ctxt
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
    87
  end
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
    88
55260
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
    89
fun preplay_trace ctxt assmsp concl outcome =
51879
ee9562d31778 added preplay tracing
smolkas
parents: 51876
diff changeset
    90
  let
ee9562d31778 added preplay tracing
smolkas
parents: 51876
diff changeset
    91
    val ctxt = ctxt |> Config.put show_markup true
55194
daa64e603e70 got rid of one of two Metis variants
blanchet
parents: 55193
diff changeset
    92
    val assms = op @ assmsp
55260
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
    93
    val time = Pretty.str ("[" ^ string_of_play_outcome outcome ^ "]")
61268
abe08fb15a12 moved remaining display.ML to more_thm.ML;
wenzelm
parents: 60549
diff changeset
    94
    val assms = Pretty.enum " and " "using " " shows " (map (Thm.pretty_thm ctxt) assms)
55251
85f5d7da4ab6 tuned code
blanchet
parents: 55244
diff changeset
    95
    val concl = Syntax.pretty_term ctxt concl
54761
0ef52f40d419 use consistent condition for setting 'metis_new_skolem' (in preplaying and in output printing) + tuning
blanchet
parents: 54712
diff changeset
    96
  in
55251
85f5d7da4ab6 tuned code
blanchet
parents: 55244
diff changeset
    97
    tracing (Pretty.string_of (Pretty.blk (2, Pretty.breaks [time, assms, concl])))
54761
0ef52f40d419 use consistent condition for setting 'metis_new_skolem' (in preplaying and in output printing) + tuning
blanchet
parents: 54712
diff changeset
    98
  end
51879
ee9562d31778 added preplay tracing
smolkas
parents: 51876
diff changeset
    99
50923
141d8f575f6f move preplaying to own structure
smolkas
parents:
diff changeset
   100
fun take_time timeout tac arg =
54761
0ef52f40d419 use consistent condition for setting 'metis_new_skolem' (in preplaying and in output printing) + tuning
blanchet
parents: 54712
diff changeset
   101
  let val timing = Timing.start () in
62519
a564458f94db tuned signature -- clarified modules;
wenzelm
parents: 62219
diff changeset
   102
    (Timeout.apply timeout tac arg; Played (#cpu (Timing.result timing)))
a564458f94db tuned signature -- clarified modules;
wenzelm
parents: 62219
diff changeset
   103
    handle Timeout.TIMEOUT _ => Play_Timed_Out timeout
50923
141d8f575f6f move preplaying to own structure
smolkas
parents:
diff changeset
   104
  end
141d8f575f6f move preplaying to own structure
smolkas
parents:
diff changeset
   105
141d8f575f6f move preplaying to own structure
smolkas
parents:
diff changeset
   106
fun resolve_fact_names ctxt names =
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
   107
  (names
54761
0ef52f40d419 use consistent condition for setting 'metis_new_skolem' (in preplaying and in output printing) + tuning
blanchet
parents: 54712
diff changeset
   108
   |>> map string_of_label
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58426
diff changeset
   109
   |> apply2 (maps (thms_of_name ctxt)))
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
  handle ERROR msg => error ("preplay error: " ^ msg)
50923
141d8f575f6f move preplaying to own structure
smolkas
parents:
diff changeset
   111
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54504
diff changeset
   112
fun thm_of_proof ctxt (Proof (fixed_frees, assms, steps)) =
50923
141d8f575f6f move preplaying to own structure
smolkas
parents:
diff changeset
   113
  let
54761
0ef52f40d419 use consistent condition for setting 'metis_new_skolem' (in preplaying and in output printing) + tuning
blanchet
parents: 54712
diff changeset
   114
    val thy = Proof_Context.theory_of ctxt
0ef52f40d419 use consistent condition for setting 'metis_new_skolem' (in preplaying and in output printing) + tuning
blanchet
parents: 54712
diff changeset
   115
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54504
diff changeset
   116
    val concl = 
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54504
diff changeset
   117
      (case try List.last steps of
55299
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55295
diff changeset
   118
        SOME (Prove (_, [], _, t, _, _, _, _)) => t
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54504
diff changeset
   119
      | _ => raise Fail "preplay error: malformed subproof")
54761
0ef52f40d419 use consistent condition for setting 'metis_new_skolem' (in preplaying and in output printing) + tuning
blanchet
parents: 54712
diff changeset
   120
51178
06689dbfe072 simplified byline, isar_qualifier
smolkas
parents: 51155
diff changeset
   121
    val var_idx = maxidx_of_term concl + 1
54761
0ef52f40d419 use consistent condition for setting 'metis_new_skolem' (in preplaying and in output printing) + tuning
blanchet
parents: 54712
diff changeset
   122
    fun var_of_free (x, T) = Var ((x, var_idx), T)
0ef52f40d419 use consistent condition for setting 'metis_new_skolem' (in preplaying and in output printing) + tuning
blanchet
parents: 54712
diff changeset
   123
    val subst = map (`var_of_free #> swap #> apfst Free) fixed_frees
51178
06689dbfe072 simplified byline, isar_qualifier
smolkas
parents: 51155
diff changeset
   124
  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
   125
    Logic.list_implies (assms |> map snd, concl)
54761
0ef52f40d419 use consistent condition for setting 'metis_new_skolem' (in preplaying and in output printing) + tuning
blanchet
parents: 54712
diff changeset
   126
    |> subst_free subst
0ef52f40d419 use consistent condition for setting 'metis_new_skolem' (in preplaying and in output printing) + tuning
blanchet
parents: 54712
diff changeset
   127
    |> Skip_Proof.make_thm thy
51178
06689dbfe072 simplified byline, isar_qualifier
smolkas
parents: 51155
diff changeset
   128
  end
06689dbfe072 simplified byline, isar_qualifier
smolkas
parents: 51155
diff changeset
   129
55299
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55295
diff changeset
   130
(* may throw exceptions *)
62219
dbac573b27e7 preplaying of 'smt' and 'metis' more in sync with actual method
blanchet
parents: 61268
diff changeset
   131
fun raw_preplay_step_for_method ctxt chained timeout meth
dbac573b27e7 preplaying of 'smt' and 'metis' more in sync with actual method
blanchet
parents: 61268
diff changeset
   132
    (Prove (_, xs, _, t, subproofs, facts, _, _)) =
55258
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   133
  let
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   134
    val goal =
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   135
      (case xs of
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   136
        [] => t
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   137
      | _ =>
60549
e168d5c48a95 keep 'Pure.all' in goals when preplaying
blanchet
parents: 59058
diff changeset
   138
        (* proof obligation: !!thesis. (!!x1...xn. t ==> thesis) ==> thesis
55258
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   139
           (cf. "~~/src/Pure/Isar/obtain.ML") *)
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   140
        let
55294
6f77310a0907 proper fresh name generation
blanchet
parents: 55287
diff changeset
   141
          val frees = map Free xs
6f77310a0907 proper fresh name generation
blanchet
parents: 55287
diff changeset
   142
          val thesis =
6f77310a0907 proper fresh name generation
blanchet
parents: 55287
diff changeset
   143
            Free (singleton (Variable.variant_frees ctxt frees) ("thesis", HOLogic.boolT))
55258
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   144
          val thesis_prop = HOLogic.mk_Trueprop thesis
50923
141d8f575f6f move preplaying to own structure
smolkas
parents:
diff changeset
   145
60549
e168d5c48a95 keep 'Pure.all' in goals when preplaying
blanchet
parents: 59058
diff changeset
   146
          (* !!x1...xn. t ==> thesis *)
55258
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   147
          val inner_prop = fold_rev Logic.all frees (Logic.mk_implies (t, thesis_prop))
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   148
        in
60549
e168d5c48a95 keep 'Pure.all' in goals when preplaying
blanchet
parents: 59058
diff changeset
   149
          (* !!thesis. (!!x1...xn. t ==> thesis) ==> thesis *)
55258
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   150
          Logic.all thesis (Logic.mk_implies (inner_prop, thesis_prop))
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   151
        end)
54761
0ef52f40d419 use consistent condition for setting 'metis_new_skolem' (in preplaying and in output printing) + tuning
blanchet
parents: 54712
diff changeset
   152
55299
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55295
diff changeset
   153
    val assmsp =
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55295
diff changeset
   154
      resolve_fact_names ctxt facts
55258
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   155
      |>> append (map (thm_of_proof ctxt) subproofs)
62219
dbac573b27e7 preplaying of 'smt' and 'metis' more in sync with actual method
blanchet
parents: 61268
diff changeset
   156
      |>> append chained
55194
daa64e603e70 got rid of one of two Metis variants
blanchet
parents: 55193
diff changeset
   157
55258
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   158
    fun prove () =
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   159
      Goal.prove ctxt [] [] goal (fn {context = ctxt, ...} =>
57054
blanchet
parents: 56093
diff changeset
   160
        HEADGOAL (tac_of_proof_method ctxt assmsp meth))
55258
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   161
      handle ERROR msg => error ("Preplay error: " ^ msg)
54761
0ef52f40d419 use consistent condition for setting 'metis_new_skolem' (in preplaying and in output printing) + tuning
blanchet
parents: 54712
diff changeset
   162
55258
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   163
    val play_outcome = take_time timeout prove ()
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   164
  in
55329
3c2dbd2e221f more generous Isar proof compression -- try to remove failing steps
blanchet
parents: 55328
diff changeset
   165
    if Config.get ctxt trace then preplay_trace ctxt assmsp goal play_outcome else ();
3c2dbd2e221f more generous Isar proof compression -- try to remove failing steps
blanchet
parents: 55328
diff changeset
   166
    play_outcome
55258
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   167
  end
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   168
62219
dbac573b27e7 preplaying of 'smt' and 'metis' more in sync with actual method
blanchet
parents: 61268
diff changeset
   169
fun preplay_isar_step_for_method ctxt chained timeout meth =
dbac573b27e7 preplaying of 'smt' and 'metis' more in sync with actual method
blanchet
parents: 61268
diff changeset
   170
  try (raw_preplay_step_for_method ctxt chained timeout meth) #> the_default Play_Failed
55278
73372494fd80 more flexible compression, choosing whichever proof method works
blanchet
parents: 55275
diff changeset
   171
57775
40eea08c0cc5 slightly earlier exit from preplaying
blanchet
parents: 57734
diff changeset
   172
val min_preplay_timeout = seconds 0.002
40eea08c0cc5 slightly earlier exit from preplaying
blanchet
parents: 57734
diff changeset
   173
62219
dbac573b27e7 preplaying of 'smt' and 'metis' more in sync with actual method
blanchet
parents: 61268
diff changeset
   174
fun preplay_isar_step ctxt chained timeout0 hopeless step =
55278
73372494fd80 more flexible compression, choosing whichever proof method works
blanchet
parents: 55275
diff changeset
   175
  let
62219
dbac573b27e7 preplaying of 'smt' and 'metis' more in sync with actual method
blanchet
parents: 61268
diff changeset
   176
    fun preplay timeout meth = (meth, preplay_isar_step_for_method ctxt chained timeout meth step)
57725
a297879fe5b8 cascading timeout in parallel evaluation, to rapidly find optimum
blanchet
parents: 57054
diff changeset
   177
    fun get_time (_, Played time) = SOME time
a297879fe5b8 cascading timeout in parallel evaluation, to rapidly find optimum
blanchet
parents: 57054
diff changeset
   178
      | get_time _ = NONE
55328
0e17e92248ac tweaked handling of 'hopeless' methods
blanchet
parents: 55314
diff changeset
   179
0e17e92248ac tweaked handling of 'hopeless' methods
blanchet
parents: 55314
diff changeset
   180
    val meths = proof_methods_of_isar_step step |> subtract (op =) hopeless
55278
73372494fd80 more flexible compression, choosing whichever proof method works
blanchet
parents: 55275
diff changeset
   181
  in
57775
40eea08c0cc5 slightly earlier exit from preplaying
blanchet
parents: 57734
diff changeset
   182
    par_list_map_filter_with_timeout get_time min_preplay_timeout timeout0 preplay meths
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58426
diff changeset
   183
    |> sort (play_outcome_ord o apply2 snd)
55278
73372494fd80 more flexible compression, choosing whichever proof method works
blanchet
parents: 55275
diff changeset
   184
  end
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   185
55260
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
   186
type isar_preplay_data = (proof_method * play_outcome Lazy.lazy) list Canonical_Label_Tab.table
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   187
55256
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
   188
fun time_of_play (Played time) = time
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
   189
  | time_of_play (Play_Timed_Out time) = time
50923
141d8f575f6f move preplaying to own structure
smolkas
parents:
diff changeset
   190
55260
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
   191
fun add_preplay_outcomes Play_Failed _ = Play_Failed
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
   192
  | add_preplay_outcomes _ Play_Failed = Play_Failed
62826
eb94e570c1a4 prefer infix operations;
wenzelm
parents: 62519
diff changeset
   193
  | add_preplay_outcomes (Played time1) (Played time2) = Played (time1 + time2)
55260
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
   194
  | add_preplay_outcomes play1 play2 =
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58426
diff changeset
   195
    Play_Timed_Out (Time.+ (apply2 time_of_play (play1, play2)))
54827
14e2c7616209 more data structure refactoring
blanchet
parents: 54826
diff changeset
   196
57054
blanchet
parents: 56093
diff changeset
   197
fun set_preplay_outcomes_of_isar_step ctxt timeout preplay_data
55299
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55295
diff changeset
   198
      (step as Prove (_, _, l, _, _, _, meths, _)) meths_outcomes0 =
55264
43473497fb65 centralize preplaying
blanchet
parents: 55260
diff changeset
   199
    let
55278
73372494fd80 more flexible compression, choosing whichever proof method works
blanchet
parents: 55275
diff changeset
   200
      fun lazy_preplay meth =
62219
dbac573b27e7 preplaying of 'smt' and 'metis' more in sync with actual method
blanchet
parents: 61268
diff changeset
   201
        Lazy.lazy (fn () => preplay_isar_step_for_method ctxt [] timeout meth step)
55278
73372494fd80 more flexible compression, choosing whichever proof method works
blanchet
parents: 55275
diff changeset
   202
      val meths_outcomes = meths_outcomes0
73372494fd80 more flexible compression, choosing whichever proof method works
blanchet
parents: 55275
diff changeset
   203
        |> map (apsnd Lazy.value)
73372494fd80 more flexible compression, choosing whichever proof method works
blanchet
parents: 55275
diff changeset
   204
        |> fold (fn meth => AList.default (op =) (meth, lazy_preplay meth)) meths
55264
43473497fb65 centralize preplaying
blanchet
parents: 55260
diff changeset
   205
    in
55308
dc68f6fb88d2 properly overwrite replay data from one compression iteration to another
blanchet
parents: 55299
diff changeset
   206
      preplay_data := Canonical_Label_Tab.update (l, fold (AList.update (op =)) meths_outcomes [])
dc68f6fb88d2 properly overwrite replay data from one compression iteration to another
blanchet
parents: 55299
diff changeset
   207
        (!preplay_data)
55264
43473497fb65 centralize preplaying
blanchet
parents: 55260
diff changeset
   208
    end
57054
blanchet
parents: 56093
diff changeset
   209
  | set_preplay_outcomes_of_isar_step _ _ _ _ _ = ()
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   210
55275
e1bf9f0c5420 less aggressive evaluation
blanchet
parents: 55272
diff changeset
   211
fun get_best_method_outcome force =
57725
a297879fe5b8 cascading timeout in parallel evaluation, to rapidly find optimum
blanchet
parents: 57054
diff changeset
   212
  tap (List.app (K () o Lazy.future Future.default_params o snd)) (* could be parallelized *)
55275
e1bf9f0c5420 less aggressive evaluation
blanchet
parents: 55272
diff changeset
   213
  #> map (apsnd force)
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58426
diff changeset
   214
  #> sort (play_outcome_ord o apply2 snd)
55275
e1bf9f0c5420 less aggressive evaluation
blanchet
parents: 55272
diff changeset
   215
  #> hd
e1bf9f0c5420 less aggressive evaluation
blanchet
parents: 55272
diff changeset
   216
55272
blanchet
parents: 55269
diff changeset
   217
fun forced_intermediate_preplay_outcome_of_isar_step preplay_data l =
55269
aae87746f412 more thorough, hybrid compression
blanchet
parents: 55268
diff changeset
   218
  let
55309
455a7f9924df don't lose additional outcomes
blanchet
parents: 55308
diff changeset
   219
    val meths_outcomes as (_, outcome1) :: _ = the (Canonical_Label_Tab.lookup preplay_data l)
55269
aae87746f412 more thorough, hybrid compression
blanchet
parents: 55268
diff changeset
   220
  in
55275
e1bf9f0c5420 less aggressive evaluation
blanchet
parents: 55272
diff changeset
   221
    if forall (not o Lazy.is_finished o snd) meths_outcomes then
e1bf9f0c5420 less aggressive evaluation
blanchet
parents: 55272
diff changeset
   222
      (case Lazy.force outcome1 of
e1bf9f0c5420 less aggressive evaluation
blanchet
parents: 55272
diff changeset
   223
        outcome as Played _ => outcome
e1bf9f0c5420 less aggressive evaluation
blanchet
parents: 55272
diff changeset
   224
      | _ => snd (get_best_method_outcome Lazy.force meths_outcomes))
e1bf9f0c5420 less aggressive evaluation
blanchet
parents: 55272
diff changeset
   225
    else
56093
4eeb73a1feec simplified preplaying information
blanchet
parents: 55452
diff changeset
   226
      let val outcome = snd (get_best_method_outcome peek_at_outcome meths_outcomes) in
4eeb73a1feec simplified preplaying information
blanchet
parents: 55452
diff changeset
   227
        if outcome = Play_Timed_Out Time.zeroTime then
4eeb73a1feec simplified preplaying information
blanchet
parents: 55452
diff changeset
   228
          snd (get_best_method_outcome Lazy.force meths_outcomes)
4eeb73a1feec simplified preplaying information
blanchet
parents: 55452
diff changeset
   229
        else
4eeb73a1feec simplified preplaying information
blanchet
parents: 55452
diff changeset
   230
          outcome
4eeb73a1feec simplified preplaying information
blanchet
parents: 55452
diff changeset
   231
      end
55269
aae87746f412 more thorough, hybrid compression
blanchet
parents: 55268
diff changeset
   232
  end
aae87746f412 more thorough, hybrid compression
blanchet
parents: 55268
diff changeset
   233
55268
blanchet
parents: 55266
diff changeset
   234
fun preplay_outcome_of_isar_step_for_method preplay_data l =
55329
3c2dbd2e221f more generous Isar proof compression -- try to remove failing steps
blanchet
parents: 55328
diff changeset
   235
  AList.lookup (op =) (the (Canonical_Label_Tab.lookup preplay_data l))
56093
4eeb73a1feec simplified preplaying information
blanchet
parents: 55452
diff changeset
   236
  #> the_default (Lazy.value (Play_Timed_Out Time.zeroTime))
55252
0dc4993b4f56 refactor data structure (step 1)
blanchet
parents: 55251
diff changeset
   237
55268
blanchet
parents: 55266
diff changeset
   238
fun fastest_method_of_isar_step preplay_data =
55269
aae87746f412 more thorough, hybrid compression
blanchet
parents: 55268
diff changeset
   239
  the o Canonical_Label_Tab.lookup preplay_data
55275
e1bf9f0c5420 less aggressive evaluation
blanchet
parents: 55272
diff changeset
   240
  #> get_best_method_outcome Lazy.force
e1bf9f0c5420 less aggressive evaluation
blanchet
parents: 55272
diff changeset
   241
  #> fst
55266
d9d31354834e centralize more preplaying
blanchet
parents: 55264
diff changeset
   242
55299
c3bb1cffce26 generate comments in Isar proofs
blanchet
parents: 55295
diff changeset
   243
fun forced_outcome_of_step preplay_data (Prove (_, _, l, _, _, _, meths, _)) =
55295
b18f65f77fcd keep all proof methods in data structure until the end, to enhance debugging output
blanchet
parents: 55294
diff changeset
   244
    Lazy.force (preplay_outcome_of_isar_step_for_method preplay_data l (hd meths))
55260
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
   245
  | forced_outcome_of_step _ _ = Played Time.zeroTime
55252
0dc4993b4f56 refactor data structure (step 1)
blanchet
parents: 55251
diff changeset
   246
55260
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
   247
fun preplay_outcome_of_isar_proof preplay_data (Proof (_, _, steps)) =
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
   248
  fold_isar_steps (add_preplay_outcomes o forced_outcome_of_step preplay_data) steps
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
   249
    (Played Time.zeroTime)
50923
141d8f575f6f move preplaying to own structure
smolkas
parents:
diff changeset
   250
54504
blanchet
parents: 53761
diff changeset
   251
end;