src/HOL/Tools/Sledgehammer/sledgehammer_isar_preplay.ML
author blanchet
Mon, 03 Feb 2014 10:14:18 +0100
changeset 55269 aae87746f412
parent 55268 a46458d368d5
child 55272 236114c5eb44
permissions -rw-r--r--
more thorough, hybrid compression
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
54828
b2271ad695db don't do 'isar_try0' if preplaying is off
blanchet
parents: 54827
diff changeset
    10
  type play_outcome = Sledgehammer_Reconstructor.play_outcome
55223
3c593bad6b31 generalized preplaying infrastructure to store various results for various methods
blanchet
parents: 55221
diff changeset
    11
  type proof_method = Sledgehammer_Isar_Proof.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
55256
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
    20
  val enrich_context_with_local_facts : isar_proof -> Proof.context -> Proof.context
55258
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
    21
  val preplay_isar_step : Proof.context -> Time.time -> proof_method -> isar_step -> play_outcome
55264
43473497fb65 centralize preplaying
blanchet
parents: 55260
diff changeset
    22
  val set_preplay_outcomes_of_isar_step : Proof.context -> Time.time ->
43473497fb65 centralize preplaying
blanchet
parents: 55260
diff changeset
    23
    isar_preplay_data Unsynchronized.ref -> isar_step ->
55260
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
    24
    (proof_method * play_outcome Lazy.lazy) list -> unit
55269
aae87746f412 more thorough, hybrid compression
blanchet
parents: 55268
diff changeset
    25
  val forced_preplay_outcome_of_isar_step : isar_preplay_data -> label -> play_outcome
55266
d9d31354834e centralize more preplaying
blanchet
parents: 55264
diff changeset
    26
  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
    27
    play_outcome Lazy.lazy
55266
d9d31354834e centralize more preplaying
blanchet
parents: 55264
diff changeset
    28
  val fastest_method_of_isar_step : isar_preplay_data -> label -> proof_method
55260
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
    29
  val preplay_outcome_of_isar_proof : isar_preplay_data -> isar_proof -> play_outcome
54504
blanchet
parents: 53761
diff changeset
    30
end;
50923
141d8f575f6f move preplaying to own structure
smolkas
parents:
diff changeset
    31
55202
824c48a539c9 renamed many Sledgehammer ML files to clarify structure
blanchet
parents: 55194
diff changeset
    32
structure Sledgehammer_Isar_Preplay : SLEDGEHAMMER_ISAR_PREPLAY =
50923
141d8f575f6f move preplaying to own structure
smolkas
parents:
diff changeset
    33
struct
141d8f575f6f move preplaying to own structure
smolkas
parents:
diff changeset
    34
55257
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55256
diff changeset
    35
open ATP_Proof_Reconstruct
50923
141d8f575f6f move preplaying to own structure
smolkas
parents:
diff changeset
    36
open Sledgehammer_Util
54828
b2271ad695db don't do 'isar_try0' if preplaying is off
blanchet
parents: 54827
diff changeset
    37
open Sledgehammer_Reconstructor
55202
824c48a539c9 renamed many Sledgehammer ML files to clarify structure
blanchet
parents: 55194
diff changeset
    38
open Sledgehammer_Isar_Proof
50923
141d8f575f6f move preplaying to own structure
smolkas
parents:
diff changeset
    39
54763
blanchet
parents: 54761
diff changeset
    40
val trace = Attrib.setup_config_bool @{binding sledgehammer_preplay_trace} (K false)
50924
beb95bf66b21 changed type of preplay time; tuned preplaying
smolkas
parents: 50923
diff changeset
    41
55256
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
    42
fun enrich_context_with_local_facts proof ctxt =
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
    43
  let
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
    44
    val thy = Proof_Context.theory_of ctxt
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
    45
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
    46
    fun enrich_with_fact l t =
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
    47
      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
    48
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
    49
    val enrich_with_assms = fold (uncurry enrich_with_fact)
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
    50
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
    51
    fun enrich_with_proof (Proof (_, assms, isar_steps)) =
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
    52
      enrich_with_assms assms #> fold enrich_with_step isar_steps
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
    53
    and enrich_with_step (Let _) = I
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
    54
      | enrich_with_step (Prove (_, _, l, t, subproofs, _)) =
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
    55
        enrich_with_fact l t #> fold enrich_with_proof subproofs
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
    56
  in
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
    57
    enrich_with_proof proof ctxt
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
    58
  end
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
    59
55260
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
    60
fun preplay_trace ctxt assmsp concl outcome =
51879
ee9562d31778 added preplay tracing
smolkas
parents: 51876
diff changeset
    61
  let
ee9562d31778 added preplay tracing
smolkas
parents: 51876
diff changeset
    62
    val ctxt = ctxt |> Config.put show_markup true
55194
daa64e603e70 got rid of one of two Metis variants
blanchet
parents: 55193
diff changeset
    63
    val assms = op @ assmsp
55260
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
    64
    val time = Pretty.str ("[" ^ string_of_play_outcome outcome ^ "]")
55251
85f5d7da4ab6 tuned code
blanchet
parents: 55244
diff changeset
    65
    val assms = Pretty.enum " and " "using " " shows " (map (Display.pretty_thm ctxt) assms)
85f5d7da4ab6 tuned code
blanchet
parents: 55244
diff changeset
    66
    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
    67
  in
55251
85f5d7da4ab6 tuned code
blanchet
parents: 55244
diff changeset
    68
    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
    69
  end
51879
ee9562d31778 added preplay tracing
smolkas
parents: 51876
diff changeset
    70
50923
141d8f575f6f move preplaying to own structure
smolkas
parents:
diff changeset
    71
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
    72
  let val timing = Timing.start () in
54828
b2271ad695db don't do 'isar_try0' if preplaying is off
blanchet
parents: 54827
diff changeset
    73
    (TimeLimit.timeLimit timeout tac arg; Played (#cpu (Timing.result timing)))
b2271ad695db don't do 'isar_try0' if preplaying is off
blanchet
parents: 54827
diff changeset
    74
    handle TimeLimit.TimeOut => Play_Timed_Out timeout
50923
141d8f575f6f move preplaying to own structure
smolkas
parents:
diff changeset
    75
  end
141d8f575f6f move preplaying to own structure
smolkas
parents:
diff changeset
    76
141d8f575f6f move preplaying to own structure
smolkas
parents:
diff changeset
    77
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
    78
  (names
54761
0ef52f40d419 use consistent condition for setting 'metis_new_skolem' (in preplaying and in output printing) + tuning
blanchet
parents: 54712
diff changeset
    79
   |>> map string_of_label
55194
daa64e603e70 got rid of one of two Metis variants
blanchet
parents: 55193
diff changeset
    80
   |> pairself (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
    81
  handle ERROR msg => error ("preplay error: " ^ msg)
50923
141d8f575f6f move preplaying to own structure
smolkas
parents:
diff changeset
    82
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54504
diff changeset
    83
fun thm_of_proof ctxt (Proof (fixed_frees, assms, steps)) =
50923
141d8f575f6f move preplaying to own structure
smolkas
parents:
diff changeset
    84
  let
54761
0ef52f40d419 use consistent condition for setting 'metis_new_skolem' (in preplaying and in output printing) + tuning
blanchet
parents: 54712
diff changeset
    85
    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
    86
54700
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54504
diff changeset
    87
    val concl = 
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54504
diff changeset
    88
      (case try List.last steps of
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54504
diff changeset
    89
        SOME (Prove (_, [], _, t, _, _)) => t
64177ce0a7bd adapted code for Z3 proof reconstruction
blanchet
parents: 54504
diff changeset
    90
      | _ => 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
    91
51178
06689dbfe072 simplified byline, isar_qualifier
smolkas
parents: 51155
diff changeset
    92
    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
    93
    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
    94
    val subst = map (`var_of_free #> swap #> apfst Free) fixed_frees
51178
06689dbfe072 simplified byline, isar_qualifier
smolkas
parents: 51155
diff changeset
    95
  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
    96
    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
    97
    |> subst_free subst
0ef52f40d419 use consistent condition for setting 'metis_new_skolem' (in preplaying and in output printing) + tuning
blanchet
parents: 54712
diff changeset
    98
    |> Skip_Proof.make_thm thy
51178
06689dbfe072 simplified byline, isar_qualifier
smolkas
parents: 51155
diff changeset
    99
  end
06689dbfe072 simplified byline, isar_qualifier
smolkas
parents: 51155
diff changeset
   100
55257
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55256
diff changeset
   101
fun tac_of_method ctxt (local_facts, global_facts) meth =
55194
daa64e603e70 got rid of one of two Metis variants
blanchet
parents: 55193
diff changeset
   102
  Method.insert_tac local_facts THEN'
54766
6ac273f176cd store alternative proof methods in Isar data structure
blanchet
parents: 54765
diff changeset
   103
  (case meth of
55194
daa64e603e70 got rid of one of two Metis variants
blanchet
parents: 55193
diff changeset
   104
    Meson_Method => Meson.meson_tac ctxt global_facts
55257
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55256
diff changeset
   105
  | Metis_Method (type_enc_opt, lam_trans_opt) =>
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55256
diff changeset
   106
    Metis_Tactic.metis_tac [type_enc_opt |> the_default partial_type_enc]
abfd7b90bba2 rationalized threading of 'metis' arguments
blanchet
parents: 55256
diff changeset
   107
      (the_default default_metis_lam_trans lam_trans_opt) ctxt global_facts
52592
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52575
diff changeset
   108
  | _ =>
55194
daa64e603e70 got rid of one of two Metis variants
blanchet
parents: 55193
diff changeset
   109
    Method.insert_tac global_facts THEN'
54766
6ac273f176cd store alternative proof methods in Isar data structure
blanchet
parents: 54765
diff changeset
   110
    (case meth of
55194
daa64e603e70 got rid of one of two Metis variants
blanchet
parents: 55193
diff changeset
   111
      Simp_Method => Simplifier.asm_full_simp_tac ctxt
54838
16511f84913c reconstruct SPASS-Pirate steps of the form 'x ~= C x' (or more complicated)
blanchet
parents: 54831
diff changeset
   112
    | Simp_Size_Method =>
16511f84913c reconstruct SPASS-Pirate steps of the form 'x ~= C x' (or more complicated)
blanchet
parents: 54831
diff changeset
   113
      Simplifier.asm_full_simp_tac (Simplifier.add_simp @{thm size_ne_size_imp_ne} ctxt)
54765
blanchet
parents: 54764
diff changeset
   114
    | Auto_Method => K (Clasimp.auto_tac ctxt)
blanchet
parents: 54764
diff changeset
   115
    | Fastforce_Method => Clasimp.fast_force_tac ctxt
blanchet
parents: 54764
diff changeset
   116
    | Force_Method => Clasimp.force_tac ctxt
blanchet
parents: 54764
diff changeset
   117
    | Arith_Method => Arith_Data.arith_tac ctxt
blanchet
parents: 54764
diff changeset
   118
    | Blast_Method => blast_tac ctxt
55219
6fe9fd75f9d7 added 'algebra' to the mix
blanchet
parents: 55213
diff changeset
   119
    | Algebra_Method => Groebner.algebra_tac [] [] ctxt
55202
824c48a539c9 renamed many Sledgehammer ML files to clarify structure
blanchet
parents: 55194
diff changeset
   120
    | _ => raise Fail "Sledgehammer_Isar_Preplay: tac_of_method"))
52592
8a25b17e3d79 optimize isar-proofs by trying different proof methods
smolkas
parents: 52575
diff changeset
   121
54761
0ef52f40d419 use consistent condition for setting 'metis_new_skolem' (in preplaying and in output printing) + tuning
blanchet
parents: 54712
diff changeset
   122
(* main function for preplaying Isar steps; may throw exceptions *)
55258
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   123
fun raw_preplay_step ctxt timeout meth (Prove (_, xs, _, t, subproofs, (fact_names, _))) =
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   124
  let
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   125
    val goal =
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   126
      (case xs of
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   127
        [] => t
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   128
      | _ =>
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   129
        (* proof obligation: !!thesis. (!!x. A x ==> thesis) ==> thesis
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   130
           (cf. "~~/src/Pure/Isar/obtain.ML") *)
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   131
        let
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   132
          (* FIXME: generate fresh name *)
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   133
          val thesis = Free ("thesis_preplay", HOLogic.boolT)
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   134
          val thesis_prop = HOLogic.mk_Trueprop thesis
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   135
          val frees = map Free xs
50923
141d8f575f6f move preplaying to own structure
smolkas
parents:
diff changeset
   136
55258
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   137
          (* !!x1..xn. t ==> thesis (xs = [x1, .., xn]) *)
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   138
          val inner_prop = fold_rev Logic.all frees (Logic.mk_implies (t, thesis_prop))
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   139
        in
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   140
          (* !!thesis. (!!x1..xn. t ==> thesis) ==> thesis *)
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   141
          Logic.all thesis (Logic.mk_implies (inner_prop, thesis_prop))
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   142
        end)
54761
0ef52f40d419 use consistent condition for setting 'metis_new_skolem' (in preplaying and in output printing) + tuning
blanchet
parents: 54712
diff changeset
   143
55258
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   144
    val facts =
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   145
      resolve_fact_names ctxt fact_names
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   146
      |>> append (map (thm_of_proof ctxt) subproofs)
55194
daa64e603e70 got rid of one of two Metis variants
blanchet
parents: 55193
diff changeset
   147
55258
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   148
    fun prove () =
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   149
      Goal.prove ctxt [] [] goal (fn {context = ctxt, ...} =>
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   150
        HEADGOAL (tac_of_method ctxt facts meth))
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   151
      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
   152
55258
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   153
    val play_outcome = take_time timeout prove ()
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   154
  in
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   155
    (if Config.get ctxt trace then preplay_trace ctxt facts goal play_outcome else ();
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   156
     play_outcome)
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   157
  end
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   158
8cc42c1f9bb5 more data structure rationalization
blanchet
parents: 55257
diff changeset
   159
fun preplay_isar_step ctxt timeout meth =
55260
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
   160
  try (raw_preplay_step ctxt timeout meth) #> the_default Play_Failed
52556
c8357085217c completely rewrote SH compress; added two parameters for experimentation/fine grained control
smolkas
parents: 52454
diff changeset
   161
55260
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
   162
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
   163
55256
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
   164
fun time_of_play (Played time) = time
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
   165
  | time_of_play (Play_Timed_Out time) = time
50923
141d8f575f6f move preplaying to own structure
smolkas
parents:
diff changeset
   166
55260
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
   167
fun add_preplay_outcomes Play_Failed _ = Play_Failed
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
   168
  | add_preplay_outcomes _ Play_Failed = Play_Failed
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
   169
  | add_preplay_outcomes (Played time1) (Played time2) = Played (Time.+ (time1, time2))
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
   170
  | add_preplay_outcomes play1 play2 =
55256
6c317e374614 refactored data structure (step 3)
blanchet
parents: 55255
diff changeset
   171
    Play_Timed_Out (Time.+ (pairself time_of_play (play1, play2)))
54827
14e2c7616209 more data structure refactoring
blanchet
parents: 54826
diff changeset
   172
55264
43473497fb65 centralize preplaying
blanchet
parents: 55260
diff changeset
   173
fun set_preplay_outcomes_of_isar_step ctxt timeout preplay_data
43473497fb65 centralize preplaying
blanchet
parents: 55260
diff changeset
   174
      (step as Prove (_, _, l, _, _, (_, meths))) meths_outcomes0 =
43473497fb65 centralize preplaying
blanchet
parents: 55260
diff changeset
   175
    let
43473497fb65 centralize preplaying
blanchet
parents: 55260
diff changeset
   176
      fun preplay meth = Lazy.lazy (fn () => preplay_isar_step ctxt timeout meth step)
43473497fb65 centralize preplaying
blanchet
parents: 55260
diff changeset
   177
      val meths_outcomes =
43473497fb65 centralize preplaying
blanchet
parents: 55260
diff changeset
   178
        fold (fn meth => AList.default (op =) (meth, preplay meth)) meths meths_outcomes0
43473497fb65 centralize preplaying
blanchet
parents: 55260
diff changeset
   179
    in
43473497fb65 centralize preplaying
blanchet
parents: 55260
diff changeset
   180
      preplay_data := Canonical_Label_Tab.map_default (l, [])
43473497fb65 centralize preplaying
blanchet
parents: 55260
diff changeset
   181
        (fold (AList.update (op =)) meths_outcomes) (!preplay_data)
43473497fb65 centralize preplaying
blanchet
parents: 55260
diff changeset
   182
    end
43473497fb65 centralize preplaying
blanchet
parents: 55260
diff changeset
   183
  | 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
   184
55269
aae87746f412 more thorough, hybrid compression
blanchet
parents: 55268
diff changeset
   185
fun peek_at_outcome outcome = if Lazy.is_finished outcome then Lazy.force outcome else Not_Played
aae87746f412 more thorough, hybrid compression
blanchet
parents: 55268
diff changeset
   186
aae87746f412 more thorough, hybrid compression
blanchet
parents: 55268
diff changeset
   187
(*
aae87746f412 more thorough, hybrid compression
blanchet
parents: 55268
diff changeset
   188
*)
aae87746f412 more thorough, hybrid compression
blanchet
parents: 55268
diff changeset
   189
fun forced_preplay_outcome_of_isar_step preplay_data l =
aae87746f412 more thorough, hybrid compression
blanchet
parents: 55268
diff changeset
   190
  let
aae87746f412 more thorough, hybrid compression
blanchet
parents: 55268
diff changeset
   191
    fun get_best_outcome_available get_one =
aae87746f412 more thorough, hybrid compression
blanchet
parents: 55268
diff changeset
   192
      the (Canonical_Label_Tab.lookup preplay_data l)
aae87746f412 more thorough, hybrid compression
blanchet
parents: 55268
diff changeset
   193
      |> map (apsnd get_one)
aae87746f412 more thorough, hybrid compression
blanchet
parents: 55268
diff changeset
   194
      |> sort (play_outcome_ord o pairself snd)
aae87746f412 more thorough, hybrid compression
blanchet
parents: 55268
diff changeset
   195
      |> hd |> snd
aae87746f412 more thorough, hybrid compression
blanchet
parents: 55268
diff changeset
   196
  in
aae87746f412 more thorough, hybrid compression
blanchet
parents: 55268
diff changeset
   197
    (case get_best_outcome_available peek_at_outcome of
aae87746f412 more thorough, hybrid compression
blanchet
parents: 55268
diff changeset
   198
      Not_Played => get_best_outcome_available Lazy.force
aae87746f412 more thorough, hybrid compression
blanchet
parents: 55268
diff changeset
   199
    | outcome => outcome)
aae87746f412 more thorough, hybrid compression
blanchet
parents: 55268
diff changeset
   200
  end
aae87746f412 more thorough, hybrid compression
blanchet
parents: 55268
diff changeset
   201
55268
blanchet
parents: 55266
diff changeset
   202
fun preplay_outcome_of_isar_step_for_method preplay_data l =
blanchet
parents: 55266
diff changeset
   203
  the o AList.lookup (op =) (the (Canonical_Label_Tab.lookup preplay_data l))
55252
0dc4993b4f56 refactor data structure (step 1)
blanchet
parents: 55251
diff changeset
   204
55268
blanchet
parents: 55266
diff changeset
   205
fun fastest_method_of_isar_step preplay_data =
55269
aae87746f412 more thorough, hybrid compression
blanchet
parents: 55268
diff changeset
   206
  the o Canonical_Label_Tab.lookup preplay_data
aae87746f412 more thorough, hybrid compression
blanchet
parents: 55268
diff changeset
   207
  #> tap (List.app (K () o Lazy.future Future.default_params o snd)) (* optional parallelism *)
aae87746f412 more thorough, hybrid compression
blanchet
parents: 55268
diff changeset
   208
  #> map (apsnd Lazy.force)
aae87746f412 more thorough, hybrid compression
blanchet
parents: 55268
diff changeset
   209
  #> sort (play_outcome_ord o pairself snd)
55268
blanchet
parents: 55266
diff changeset
   210
  #> hd #> fst
55266
d9d31354834e centralize more preplaying
blanchet
parents: 55264
diff changeset
   211
d9d31354834e centralize more preplaying
blanchet
parents: 55264
diff changeset
   212
fun forced_outcome_of_step preplay_data (Prove (_, _, l, _, _, (_, meths))) =
d9d31354834e centralize more preplaying
blanchet
parents: 55264
diff changeset
   213
    Lazy.force (preplay_outcome_of_isar_step_for_method preplay_data l (the_single meths))
55260
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
   214
  | forced_outcome_of_step _ _ = Played Time.zeroTime
55252
0dc4993b4f56 refactor data structure (step 1)
blanchet
parents: 55251
diff changeset
   215
55260
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
   216
fun preplay_outcome_of_isar_proof preplay_data (Proof (_, _, steps)) =
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
   217
  fold_isar_steps (add_preplay_outcomes o forced_outcome_of_step preplay_data) steps
ada3ae6458d4 more data structure rationalization
blanchet
parents: 55258
diff changeset
   218
    (Played Time.zeroTime)
50923
141d8f575f6f move preplaying to own structure
smolkas
parents:
diff changeset
   219
54504
blanchet
parents: 53761
diff changeset
   220
end;