src/HOL/Tools/Sledgehammer/sledgehammer_filter_mash.ML
author blanchet
Wed, 18 Jul 2012 08:44:04 +0200
changeset 49321 e699f754d9f7
parent 49320 399f7e20e17f
child 49323 89674e5a4d35
permissions -rw-r--r--
better zipping of MaSh facts
blanchet@49263
     1
(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_filter_mash.ML
blanchet@49263
     2
    Author:     Jasmin Blanchette, TU Muenchen
blanchet@49263
     3
blanchet@49263
     4
Sledgehammer's machine-learning-based relevance filter (MaSh).
blanchet@49263
     5
*)
blanchet@49263
     6
blanchet@49263
     7
signature SLEDGEHAMMER_FILTER_MASH =
blanchet@49263
     8
sig
blanchet@49266
     9
  type status = ATP_Problem_Generate.status
blanchet@49266
    10
  type stature = ATP_Problem_Generate.stature
blanchet@49311
    11
  type fact = Sledgehammer_Fact.fact
blanchet@49311
    12
  type fact_override = Sledgehammer_Fact.fact_override
blanchet@49266
    13
  type params = Sledgehammer_Provers.params
blanchet@49303
    14
  type relevance_fudge = Sledgehammer_Provers.relevance_fudge
blanchet@49266
    15
  type prover_result = Sledgehammer_Provers.prover_result
blanchet@49266
    16
blanchet@49318
    17
  val escape_meta : string -> string
blanchet@49318
    18
  val escape_metas : string list -> string
blanchet@49314
    19
  val all_non_tautological_facts_of :
blanchet@49314
    20
    theory -> status Termtab.table -> fact list
blanchet@49266
    21
  val theory_ord : theory * theory -> order
blanchet@49266
    22
  val thm_ord : thm * thm -> order
blanchet@49318
    23
  val thy_facts_from_thms : ('a * thm) list -> string list Symtab.table
blanchet@49266
    24
  val has_thy : theory -> thm -> bool
blanchet@49318
    25
  val parent_facts : theory -> string list Symtab.table -> string list
blanchet@49317
    26
  val features_of : theory -> status -> term list -> string list
blanchet@49266
    27
  val isabelle_dependencies_of : string list -> thm -> string list
blanchet@49266
    28
  val goal_of_thm : theory -> thm -> thm
blanchet@49311
    29
  val run_prover : Proof.context -> params -> fact list -> thm -> prover_result
blanchet@49319
    30
  val thy_name_of_fact : string -> string
blanchet@49317
    31
  val mash_RESET : unit -> unit
blanchet@49319
    32
  val mash_ADD : (string * string list * string list * string list) list -> unit
blanchet@49317
    33
  val mash_DEL : string list -> string list -> unit
blanchet@49317
    34
  val mash_SUGGEST : string list -> string list -> string list
blanchet@49313
    35
  val mash_reset : unit -> unit
blanchet@49313
    36
  val mash_can_suggest_facts : unit -> bool
blanchet@49313
    37
  val mash_suggest_facts :
blanchet@49317
    38
    Proof.context -> params -> string -> int -> term list -> term -> fact list
blanchet@49313
    39
    -> fact list * fact list
blanchet@49313
    40
  val mash_can_learn_thy : theory -> bool
blanchet@49319
    41
  val mash_learn_thy : Proof.context -> real -> unit
blanchet@49313
    42
  val mash_learn_proof : theory -> term -> thm list -> unit
blanchet@49303
    43
  val relevant_facts :
blanchet@49307
    44
    Proof.context -> params -> string -> int -> fact_override -> term list
blanchet@49311
    45
    -> term -> fact list -> fact list
blanchet@49263
    46
end;
blanchet@49263
    47
blanchet@49263
    48
structure Sledgehammer_Filter_MaSh : SLEDGEHAMMER_FILTER_MASH =
blanchet@49263
    49
struct
blanchet@49264
    50
blanchet@49266
    51
open ATP_Util
blanchet@49266
    52
open ATP_Problem_Generate
blanchet@49266
    53
open Sledgehammer_Util
blanchet@49266
    54
open Sledgehammer_Fact
blanchet@49266
    55
open Sledgehammer_Filter_Iter
blanchet@49266
    56
open Sledgehammer_Provers
blanchet@49266
    57
blanchet@49316
    58
val mash_dir = "mash"
blanchet@49316
    59
val model_file = "model"
blanchet@49316
    60
val state_file = "state"
blanchet@49316
    61
blanchet@49317
    62
fun mk_path file =
blanchet@49317
    63
  getenv "ISABELLE_HOME_USER" ^ "/" ^ mash_dir ^ "/" ^ file
blanchet@49317
    64
  |> Path.explode
blanchet@49316
    65
blanchet@49317
    66
val fresh_fact_prefix = Long_Name.separator
blanchet@49266
    67
blanchet@49266
    68
(*** Isabelle helpers ***)
blanchet@49266
    69
blanchet@49266
    70
fun escape_meta_char c =
blanchet@49266
    71
  if Char.isAlphaNum c orelse c = #"_" orelse c = #"." orelse c = #"(" orelse
blanchet@49266
    72
     c = #")" orelse c = #"," then
blanchet@49266
    73
    String.str c
blanchet@49266
    74
  else
blanchet@49266
    75
    (* fixed width, in case more digits follow *)
blanchet@49266
    76
    "\\" ^ stringN_of_int 3 (Char.ord c)
blanchet@49266
    77
blanchet@49266
    78
val escape_meta = String.translate escape_meta_char
blanchet@49318
    79
val escape_metas = map escape_meta #> space_implode " "
blanchet@49266
    80
blanchet@49318
    81
val thy_feature_prefix = "y_"
blanchet@49266
    82
blanchet@49318
    83
val thy_feature_name_of = prefix thy_feature_prefix
blanchet@49318
    84
val const_name_of = prefix const_prefix
blanchet@49318
    85
val type_name_of = prefix type_const_prefix
blanchet@49318
    86
val class_name_of = prefix class_prefix
blanchet@49266
    87
blanchet@49266
    88
local
blanchet@49266
    89
blanchet@49266
    90
fun has_bool @{typ bool} = true
blanchet@49266
    91
  | has_bool (Type (_, Ts)) = exists has_bool Ts
blanchet@49266
    92
  | has_bool _ = false
blanchet@49266
    93
blanchet@49266
    94
fun has_fun (Type (@{type_name fun}, _)) = true
blanchet@49266
    95
  | has_fun (Type (_, Ts)) = exists has_fun Ts
blanchet@49266
    96
  | has_fun _ = false
blanchet@49266
    97
blanchet@49266
    98
val is_conn = member (op =)
blanchet@49266
    99
  [@{const_name Trueprop}, @{const_name HOL.conj}, @{const_name HOL.disj},
blanchet@49266
   100
   @{const_name HOL.implies}, @{const_name Not},
blanchet@49266
   101
   @{const_name All}, @{const_name Ex}, @{const_name Ball}, @{const_name Bex},
blanchet@49266
   102
   @{const_name HOL.eq}]
blanchet@49266
   103
blanchet@49266
   104
val has_bool_arg_const =
blanchet@49266
   105
  exists_Const (fn (c, T) =>
blanchet@49266
   106
                   not (is_conn c) andalso exists has_bool (binder_types T))
blanchet@49266
   107
blanchet@49266
   108
fun higher_inst_const thy (c, T) =
blanchet@49266
   109
  case binder_types T of
blanchet@49266
   110
    [] => false
blanchet@49266
   111
  | Ts => length (binder_types (Sign.the_const_type thy c)) <> length Ts
blanchet@49266
   112
blanchet@49266
   113
val binders = [@{const_name All}, @{const_name Ex}]
blanchet@49266
   114
blanchet@49266
   115
in
blanchet@49266
   116
blanchet@49266
   117
fun is_fo_term thy t =
blanchet@49266
   118
  let
blanchet@49266
   119
    val t =
blanchet@49266
   120
      t |> Envir.beta_eta_contract
blanchet@49266
   121
        |> transform_elim_prop
blanchet@49266
   122
        |> Object_Logic.atomize_term thy
blanchet@49266
   123
  in
blanchet@49266
   124
    Term.is_first_order binders t andalso
blanchet@49266
   125
    not (exists_subterm (fn Var (_, T) => has_bool T orelse has_fun T
blanchet@49266
   126
                          | _ => false) t orelse
blanchet@49266
   127
         has_bool_arg_const t orelse exists_Const (higher_inst_const thy) t)
blanchet@49266
   128
  end
blanchet@49266
   129
blanchet@49266
   130
end
blanchet@49266
   131
blanchet@49317
   132
fun interesting_terms_types_and_classes term_max_depth type_max_depth ts =
blanchet@49266
   133
  let
blanchet@49266
   134
    val bad_types = [@{type_name prop}, @{type_name bool}, @{type_name fun}]
blanchet@49266
   135
    val bad_consts = atp_widely_irrelevant_consts
blanchet@49319
   136
    fun add_classes @{sort type} = I
blanchet@49319
   137
      | add_classes S = union (op =) (map class_name_of S)
blanchet@49266
   138
    fun do_add_type (Type (s, Ts)) =
blanchet@49266
   139
        (not (member (op =) bad_types s) ? insert (op =) (type_name_of s))
blanchet@49266
   140
        #> fold do_add_type Ts
blanchet@49319
   141
      | do_add_type (TFree (_, S)) = add_classes S
blanchet@49319
   142
      | do_add_type (TVar (_, S)) = add_classes S
blanchet@49266
   143
    fun add_type T = type_max_depth >= 0 ? do_add_type T
blanchet@49266
   144
    fun mk_app s args =
blanchet@49266
   145
      if member (op <>) args "" then s ^ "(" ^ space_implode "," args ^ ")"
blanchet@49266
   146
      else s
blanchet@49266
   147
    fun patternify ~1 _ = ""
blanchet@49266
   148
      | patternify depth t =
blanchet@49266
   149
        case strip_comb t of
blanchet@49266
   150
          (Const (s, _), args) =>
blanchet@49266
   151
          mk_app (const_name_of s) (map (patternify (depth - 1)) args)
blanchet@49266
   152
        | _ => ""
blanchet@49266
   153
    fun add_term_patterns ~1 _ = I
blanchet@49266
   154
      | add_term_patterns depth t =
blanchet@49266
   155
        insert (op =) (patternify depth t)
blanchet@49266
   156
        #> add_term_patterns (depth - 1) t
blanchet@49266
   157
    val add_term = add_term_patterns term_max_depth
blanchet@49266
   158
    fun add_patterns t =
blanchet@49266
   159
      let val (head, args) = strip_comb t in
blanchet@49266
   160
        (case head of
blanchet@49266
   161
           Const (s, T) =>
blanchet@49266
   162
           not (member (op =) bad_consts s) ? (add_term t #> add_type T)
blanchet@49266
   163
         | Free (_, T) => add_type T
blanchet@49266
   164
         | Var (_, T) => add_type T
blanchet@49266
   165
         | Abs (_, T, body) => add_type T #> add_patterns body
blanchet@49266
   166
         | _ => I)
blanchet@49266
   167
        #> fold add_patterns args
blanchet@49266
   168
      end
blanchet@49317
   169
  in [] |> fold add_patterns ts |> sort string_ord end
blanchet@49266
   170
blanchet@49266
   171
fun is_likely_tautology th =
blanchet@49317
   172
  null (interesting_terms_types_and_classes 0 ~1 [prop_of th]) andalso
blanchet@49266
   173
  not (Thm.eq_thm_prop (@{thm ext}, th))
blanchet@49266
   174
blanchet@49266
   175
fun is_too_meta thy th =
blanchet@49266
   176
  fastype_of (Object_Logic.atomize_term thy (prop_of th)) <> @{typ bool}
blanchet@49266
   177
blanchet@49314
   178
fun all_non_tautological_facts_of thy css_table =
blanchet@49314
   179
  all_facts_of thy css_table
blanchet@49266
   180
  |> filter_out ((is_likely_tautology orf is_too_meta thy) o snd)
blanchet@49266
   181
blanchet@49266
   182
fun theory_ord p =
blanchet@49266
   183
  if Theory.eq_thy p then EQUAL
blanchet@49266
   184
  else if Theory.subthy p then LESS
blanchet@49266
   185
  else if Theory.subthy (swap p) then GREATER
blanchet@49266
   186
  else EQUAL
blanchet@49266
   187
blanchet@49266
   188
val thm_ord = theory_ord o pairself theory_of_thm
blanchet@49266
   189
blanchet@49318
   190
(* ### FIXME: optimize *)
blanchet@49318
   191
fun thy_facts_from_thms ths =
blanchet@49318
   192
  ths |> map (snd #> `(theory_of_thm #> Context.theory_name))
blanchet@49266
   193
      |> AList.group (op =)
blanchet@49266
   194
      |> sort (int_ord o pairself (length o Theory.ancestors_of o theory_of_thm
blanchet@49266
   195
                                   o hd o snd))
blanchet@49318
   196
      |> map (apsnd (sort (rev_order o thm_ord) #> map Thm.get_name_hint))
blanchet@49318
   197
      |> Symtab.make
blanchet@49266
   198
blanchet@49318
   199
fun has_thy thy th =
blanchet@49318
   200
  Context.theory_name thy = Context.theory_name (theory_of_thm th)
blanchet@49266
   201
blanchet@49318
   202
fun parent_facts thy thy_facts =
blanchet@49318
   203
  let
blanchet@49318
   204
    fun add_last thy =
blanchet@49318
   205
      case Symtab.lookup thy_facts (Context.theory_name thy) of
blanchet@49318
   206
        SOME (last_fact :: _) => insert (op =) last_fact
blanchet@49318
   207
      | _ => add_parent thy
blanchet@49318
   208
    and add_parent thy = fold add_last (Theory.parents_of thy)
blanchet@49318
   209
  in add_parent thy [] end
blanchet@49266
   210
blanchet@49266
   211
fun is_exists (s, _) = (s = @{const_name Ex} orelse s = @{const_name Ex1})
blanchet@49266
   212
blanchet@49312
   213
val term_max_depth = 1
blanchet@49312
   214
val type_max_depth = 1
blanchet@49266
   215
blanchet@49266
   216
(* TODO: Generate type classes for types? *)
blanchet@49317
   217
fun features_of thy status ts =
blanchet@49318
   218
  thy_feature_name_of (Context.theory_name thy) ::
blanchet@49317
   219
  interesting_terms_types_and_classes term_max_depth type_max_depth ts
blanchet@49317
   220
  |> exists (not o is_lambda_free) ts ? cons "lambdas"
blanchet@49317
   221
  |> exists (exists_Const is_exists) ts ? cons "skolems"
blanchet@49317
   222
  |> exists (not o is_fo_term thy) ts ? cons "ho"
blanchet@49317
   223
  |> (case status of
blanchet@49317
   224
        General => I
blanchet@49317
   225
      | Induction => cons "induction"
blanchet@49317
   226
      | Intro => cons "intro"
blanchet@49317
   227
      | Inductive => cons "inductive"
blanchet@49317
   228
      | Elim => cons "elim"
blanchet@49317
   229
      | Simp => cons "simp"
blanchet@49317
   230
      | Def => cons "def")
blanchet@49266
   231
blanchet@49266
   232
fun isabelle_dependencies_of all_facts =
blanchet@49318
   233
  thms_in_proof (SOME all_facts) #> sort string_ord
blanchet@49266
   234
blanchet@49266
   235
val freezeT = Type.legacy_freeze_type
blanchet@49266
   236
blanchet@49266
   237
fun freeze (t $ u) = freeze t $ freeze u
blanchet@49266
   238
  | freeze (Abs (s, T, t)) = Abs (s, freezeT T, freeze t)
blanchet@49266
   239
  | freeze (Var ((s, _), T)) = Free (s, freezeT T)
blanchet@49266
   240
  | freeze (Const (s, T)) = Const (s, freezeT T)
blanchet@49266
   241
  | freeze (Free (s, T)) = Free (s, freezeT T)
blanchet@49266
   242
  | freeze t = t
blanchet@49266
   243
blanchet@49266
   244
fun goal_of_thm thy = prop_of #> freeze #> cterm_of thy #> Goal.init
blanchet@49266
   245
blanchet@49266
   246
fun run_prover ctxt (params as {provers, ...}) facts goal =
blanchet@49266
   247
  let
blanchet@49266
   248
    val problem =
blanchet@49266
   249
      {state = Proof.init ctxt, goal = goal, subgoal = 1, subgoal_count = 1,
blanchet@49304
   250
       facts = facts |> map (apfst (apfst (fn name => name ())))
blanchet@49304
   251
                     |> map Sledgehammer_Provers.Untranslated_Fact}
blanchet@49266
   252
    val prover =
blanchet@49266
   253
      Sledgehammer_Minimize.get_minimizing_prover ctxt
blanchet@49266
   254
          Sledgehammer_Provers.Normal (hd provers)
blanchet@49266
   255
  in prover params (K (K (K ""))) problem end
blanchet@49266
   256
blanchet@49318
   257
(* ###
blanchet@49318
   258
fun compute_accessibility thy thy_facts =
blanchet@49318
   259
  let
blanchet@49318
   260
    fun add_facts (f :: fs) prevs = (f, prevs) :: add_facts fs [th]
blanchet@49318
   261
    fun add_thy facts =
blanchet@49318
   262
      let
blanchet@49318
   263
        val thy = theory_of_thm (hd facts)
blanchet@49318
   264
        val parents = parent_facts thy_facts thy
blanchet@49319
   265
      in add_facts facts parents end
blanchet@49318
   266
  in fold (add_thy o snd) thy_facts end
blanchet@49318
   267
*)
blanchet@49318
   268
blanchet@49318
   269
fun accessibility_of thy thy_facts =
blanchet@49318
   270
  case Symtab.lookup thy_facts (Context.theory_name thy) of
blanchet@49318
   271
    SOME (fact :: _) => [fact]
blanchet@49318
   272
  | _ => parent_facts thy thy_facts
blanchet@49318
   273
blanchet@49319
   274
val thy_name_of_fact = hd o Long_Name.explode
blanchet@49266
   275
blanchet@49266
   276
blanchet@49317
   277
(*** Low-level communication with MaSh ***)
blanchet@49317
   278
blanchet@49317
   279
fun mash_RESET () = (warning "MaSh RESET"; File.rm (mk_path model_file))
blanchet@49317
   280
blanchet@49319
   281
val mash_ADD =
blanchet@49319
   282
  let
blanchet@49320
   283
    fun add_record (fact, access, feats, deps) =
blanchet@49319
   284
      let
blanchet@49319
   285
        val s =
blanchet@49319
   286
          escape_meta fact ^ ": " ^ escape_metas access ^ "; " ^
blanchet@49319
   287
          escape_metas feats ^ "; " ^ escape_metas deps
blanchet@49319
   288
      in warning ("MaSh ADD " ^ s) end
blanchet@49320
   289
  in List.app add_record end
blanchet@49317
   290
blanchet@49317
   291
fun mash_DEL facts feats =
blanchet@49318
   292
  warning ("MaSh DEL " ^ escape_metas facts ^ "; " ^ escape_metas feats)
blanchet@49317
   293
blanchet@49317
   294
fun mash_SUGGEST access feats =
blanchet@49318
   295
  (warning ("MaSh SUGGEST " ^ escape_metas access ^ "; " ^ escape_metas feats);
blanchet@49317
   296
   [])
blanchet@49317
   297
blanchet@49317
   298
blanchet@49266
   299
(*** High-level communication with MaSh ***)
blanchet@49266
   300
blanchet@49316
   301
type mash_state =
blanchet@49317
   302
  {fresh : int,
blanchet@49321
   303
   dirty_thys : unit Symtab.table,
blanchet@49316
   304
   thy_facts : string list Symtab.table}
blanchet@49264
   305
blanchet@49319
   306
val empty_state =
blanchet@49317
   307
  {fresh = 0,
blanchet@49321
   308
   dirty_thys = Symtab.empty,
blanchet@49316
   309
   thy_facts = Symtab.empty}
blanchet@49316
   310
blanchet@49316
   311
local
blanchet@49316
   312
blanchet@49316
   313
fun mash_load (state as (true, _)) = state
blanchet@49316
   314
  | mash_load _ =
blanchet@49317
   315
    let
blanchet@49317
   316
      val path = mk_path state_file
blanchet@49317
   317
      val _ = Isabelle_System.mkdir (path |> Path.dir)
blanchet@49317
   318
    in
blanchet@49317
   319
      (true,
blanchet@49317
   320
       case try File.read_lines path of
blanchet@49321
   321
         SOME (fresh_line :: dirty_line :: facts_lines) =>
blanchet@49317
   322
         let
blanchet@49321
   323
           fun dirty_thys_of_line line =
blanchet@49321
   324
             Symtab.make (line |> space_explode " " |> map (rpair ()))
blanchet@49317
   325
           fun add_facts_line line =
blanchet@49317
   326
             case space_explode " " line of
blanchet@49317
   327
               thy :: facts => Symtab.update_new (thy, facts)
blanchet@49317
   328
             | _ => I (* shouldn't happen *)
blanchet@49317
   329
         in
blanchet@49317
   330
           {fresh = Int.fromString fresh_line |> the_default 0,
blanchet@49321
   331
            dirty_thys = dirty_thys_of_line dirty_line,
blanchet@49317
   332
            thy_facts = fold add_facts_line facts_lines Symtab.empty}
blanchet@49317
   333
         end
blanchet@49319
   334
       | _ => empty_state)
blanchet@49317
   335
    end
blanchet@49316
   336
blanchet@49321
   337
fun mash_save ({fresh, dirty_thys, thy_facts} : mash_state) =
blanchet@49316
   338
  let
blanchet@49317
   339
    val path = mk_path state_file
blanchet@49321
   340
    val dirty_line = (dirty_thys |> Symtab.keys |> escape_metas) ^ "\n"
blanchet@49318
   341
    fun fact_line_for (thy, facts) = escape_metas (thy :: facts) ^ "\n"
blanchet@49316
   342
  in
blanchet@49321
   343
    File.write path (string_of_int fresh ^ "\n" ^ dirty_line);
blanchet@49316
   344
    Symtab.fold (fn thy_fact => fn () =>
blanchet@49316
   345
                    File.append path (fact_line_for thy_fact)) thy_facts
blanchet@49316
   346
  end
blanchet@49316
   347
blanchet@49317
   348
val global_state =
blanchet@49319
   349
  Synchronized.var "Sledgehammer_Filter_MaSh.global_state" (false, empty_state)
blanchet@49316
   350
blanchet@49316
   351
in
blanchet@49316
   352
blanchet@49321
   353
fun mash_map f =
blanchet@49317
   354
  Synchronized.change global_state (mash_load ##> (f #> tap mash_save))
blanchet@49316
   355
blanchet@49317
   356
fun mash_peek f = Synchronized.change global_state (mash_load ##> tap f)
blanchet@49317
   357
blanchet@49319
   358
fun mash_get () = Synchronized.change_result global_state (mash_load #> `snd)
blanchet@49317
   359
blanchet@49317
   360
fun mash_reset () =
blanchet@49317
   361
  Synchronized.change global_state (fn _ =>
blanchet@49319
   362
      (mash_RESET (); File.rm (mk_path state_file); (true, empty_state)))
blanchet@49316
   363
blanchet@49316
   364
end
blanchet@49316
   365
blanchet@49319
   366
fun mash_can_suggest_facts () = not (Symtab.is_empty (#thy_facts (mash_get ())))
blanchet@49264
   367
blanchet@49313
   368
fun mash_suggest_facts ctxt params prover max_facts hyp_ts concl_t facts =
blanchet@49316
   369
  let
blanchet@49317
   370
    val thy = Proof_Context.theory_of ctxt
blanchet@49319
   371
    val access = accessibility_of thy (#thy_facts (mash_get ()))
blanchet@49317
   372
    val feats = features_of thy General (concl_t :: hyp_ts)
blanchet@49316
   373
    val suggs = mash_SUGGEST access feats
blanchet@49316
   374
  in (facts, []) end
blanchet@49264
   375
blanchet@49316
   376
fun mash_can_learn_thy thy =
blanchet@49321
   377
  not (Symtab.defined (#dirty_thys (mash_get ())) (Context.theory_name thy))
blanchet@49264
   378
blanchet@49319
   379
fun is_fact_in_thy_facts thy_facts fact =
blanchet@49319
   380
  case Symtab.lookup thy_facts (thy_name_of_fact fact) of
blanchet@49319
   381
    SOME facts => member (op =) facts fact
blanchet@49319
   382
  | NONE => false
blanchet@49319
   383
blanchet@49321
   384
fun zip_facts news [] = news
blanchet@49321
   385
  | zip_facts [] olds = olds
blanchet@49321
   386
  | zip_facts (new :: news) (old :: olds) =
blanchet@49321
   387
    if new = old then
blanchet@49321
   388
      new :: zip_facts news olds
blanchet@49321
   389
    else if member (op =) news old then
blanchet@49321
   390
      old :: zip_facts (filter_out (curry (op =) old) news) olds
blanchet@49321
   391
    else if member (op =) olds new then
blanchet@49321
   392
      new :: zip_facts news (filter_out (curry (op =) new) olds)
blanchet@49321
   393
    else
blanchet@49321
   394
      new :: old :: zip_facts news olds
blanchet@49321
   395
blanchet@49321
   396
fun add_thy_facts_from_thys new old =
blanchet@49321
   397
  let
blanchet@49321
   398
    fun add_thy (thy, new_facts) =
blanchet@49321
   399
      case Symtab.lookup old thy of
blanchet@49321
   400
        SOME old_facts => Symtab.update (thy, zip_facts old_facts new_facts)
blanchet@49321
   401
      | NONE => Symtab.update_new (thy, new_facts)
blanchet@49321
   402
  in old |> Symtab.fold add_thy new end
blanchet@49321
   403
blanchet@49319
   404
fun mash_learn_thy ctxt timeout =
blanchet@49319
   405
  let
blanchet@49319
   406
    val thy = Proof_Context.theory_of ctxt
blanchet@49319
   407
    val css_table = Sledgehammer_Fact.clasimpset_rule_table_of ctxt
blanchet@49319
   408
    val facts = all_non_tautological_facts_of thy css_table
blanchet@49321
   409
    val {thy_facts, ...} = mash_get ()
blanchet@49319
   410
    fun is_old (_, th) = is_fact_in_thy_facts thy_facts (Thm.get_name_hint th)
blanchet@49319
   411
    val (old_facts, new_facts) =
blanchet@49319
   412
      facts |> List.partition is_old ||> sort (thm_ord o pairself snd)
blanchet@49319
   413
    val ths = facts |> map snd
blanchet@49319
   414
    val all_names = ths |> map Thm.get_name_hint
blanchet@49320
   415
    fun do_fact ((_, (_, status)), th) (prevs, records) =
blanchet@49319
   416
      let
blanchet@49319
   417
        val name = Thm.get_name_hint th
blanchet@49319
   418
        val feats = features_of thy status [prop_of th]
blanchet@49319
   419
        val deps = isabelle_dependencies_of all_names th
blanchet@49320
   420
        val record = (name, prevs, feats, deps)
blanchet@49320
   421
      in ([name], record :: records) end
blanchet@49319
   422
    val parents = parent_facts thy thy_facts
blanchet@49320
   423
    val (_, records) = (parents, []) |> fold_rev do_fact new_facts
blanchet@49321
   424
    val new_thy_facts = new_facts |> thy_facts_from_thms
blanchet@49321
   425
    fun trans {fresh, dirty_thys, thy_facts} =
blanchet@49321
   426
      (mash_ADD records;
blanchet@49321
   427
       {fresh = fresh, dirty_thys = dirty_thys,
blanchet@49321
   428
        thy_facts = thy_facts |> add_thy_facts_from_thys new_thy_facts})
blanchet@49321
   429
  in mash_map trans end
blanchet@49319
   430
blanchet@49319
   431
(* ###
blanchet@49319
   432
fun compute_accessibility thy thy_facts =
blanchet@49319
   433
  let
blanchet@49319
   434
    fun add_facts (f :: fs) prevs = (f, prevs) :: add_facts fs [th]
blanchet@49319
   435
    fun add_thy facts =
blanchet@49319
   436
      let
blanchet@49319
   437
        val thy = theory_of_thm (hd facts)
blanchet@49319
   438
        val parents = parent_facts thy_facts thy
blanchet@49319
   439
      in add_facts facts parents end
blanchet@49319
   440
  in fold (add_thy o snd) thy_facts end
blanchet@49319
   441
*)
blanchet@49313
   442
blanchet@49317
   443
fun mash_learn_proof thy t ths =
blanchet@49321
   444
  mash_map (fn state as {fresh, dirty_thys, thy_facts} =>
blanchet@49319
   445
    let val deps = ths |> map Thm.get_name_hint in
blanchet@49319
   446
      if forall (is_fact_in_thy_facts thy_facts) deps then
blanchet@49319
   447
        let
blanchet@49319
   448
          val fact = fresh_fact_prefix ^ string_of_int fresh
blanchet@49319
   449
          val access = accessibility_of thy thy_facts
blanchet@49319
   450
          val feats = features_of thy General [t]
blanchet@49319
   451
        in
blanchet@49319
   452
          mash_ADD [(fact, access, feats, deps)];
blanchet@49321
   453
          {fresh = fresh + 1, dirty_thys = dirty_thys, thy_facts = thy_facts}
blanchet@49319
   454
        end
blanchet@49319
   455
      else
blanchet@49319
   456
        state
blanchet@49317
   457
    end)
blanchet@49264
   458
blanchet@49308
   459
fun relevant_facts ctxt params prover max_facts
blanchet@49313
   460
        ({add, only, ...} : fact_override) hyp_ts concl_t facts =
blanchet@49303
   461
  if only then
blanchet@49304
   462
    facts
blanchet@49308
   463
  else if max_facts <= 0 then
blanchet@49303
   464
    []
blanchet@49303
   465
  else
blanchet@49303
   466
    let
blanchet@49303
   467
      val add_ths = Attrib.eval_thms ctxt add
blanchet@49307
   468
      fun prepend_facts ths accepts =
blanchet@49303
   469
        ((facts |> filter (member Thm.eq_thm_prop ths o snd)) @
blanchet@49307
   470
         (accepts |> filter_out (member Thm.eq_thm_prop ths o snd)))
blanchet@49308
   471
        |> take max_facts
blanchet@49313
   472
      val iter_facts =
blanchet@49313
   473
        iterative_relevant_facts ctxt params prover max_facts NONE hyp_ts
blanchet@49313
   474
                                 concl_t facts
blanchet@49313
   475
      val (mash_facts, mash_rejected) =
blanchet@49313
   476
        mash_suggest_facts ctxt params prover max_facts hyp_ts concl_t facts
blanchet@49317
   477
      val mesh_facts = iter_facts (* ### *)
blanchet@49303
   478
    in
blanchet@49313
   479
      mesh_facts
blanchet@49303
   480
      |> not (null add_ths) ? prepend_facts add_ths
blanchet@49303
   481
    end
blanchet@49303
   482
blanchet@49263
   483
end;