src/HOL/TPTP/mash_export.ML
author blanchet
Thu, 26 Jul 2012 10:48:03 +0200
changeset 49546 7da5d3b8aef4
parent 49545 d443166f9520
child 49680 14b0732c72f7
permissions -rw-r--r--
don't export technical theorems for MaSh
blanchet@49249
     1
(*  Title:      HOL/TPTP/mash_export.ML
blanchet@49249
     2
    Author:     Jasmin Blanchette, TU Muenchen
blanchet@49249
     3
    Copyright   2012
blanchet@49249
     4
blanchet@49249
     5
Export Isabelle theory information for MaSh (Machine-learning for Sledgehammer).
blanchet@49249
     6
*)
blanchet@49249
     7
blanchet@49249
     8
signature MASH_EXPORT =
blanchet@49249
     9
sig
blanchet@49266
    10
  type params = Sledgehammer_Provers.params
blanchet@49250
    11
blanchet@49545
    12
  val generate_accessibility : Proof.context -> theory -> bool -> string -> unit
blanchet@49333
    13
  val generate_features :
blanchet@49333
    14
    Proof.context -> string -> theory -> bool -> string -> unit
blanchet@49348
    15
  val generate_isar_dependencies :
blanchet@49545
    16
    Proof.context -> theory -> bool -> string -> unit
blanchet@49348
    17
  val generate_atp_dependencies :
blanchet@49319
    18
    Proof.context -> params -> theory -> bool -> string -> unit
blanchet@49544
    19
  val generate_commands : Proof.context -> params -> theory -> string -> unit
blanchet@49394
    20
  val generate_mepo_suggestions :
blanchet@49266
    21
    Proof.context -> params -> theory -> int -> string -> unit
blanchet@49249
    22
end;
blanchet@49249
    23
blanchet@49249
    24
structure MaSh_Export : MASH_EXPORT =
blanchet@49249
    25
struct
blanchet@49249
    26
blanchet@49319
    27
open Sledgehammer_Fact
blanchet@49396
    28
open Sledgehammer_MePo
blanchet@49396
    29
open Sledgehammer_MaSh
blanchet@49260
    30
blanchet@49331
    31
fun thy_map_from_facts ths =
blanchet@49331
    32
  ths |> sort (thm_ord o swap o pairself snd)
blanchet@49331
    33
      |> map (snd #> `(theory_of_thm #> Context.theory_name))
blanchet@49331
    34
      |> AList.coalesce (op =)
blanchet@49393
    35
      |> map (apsnd (map nickname_of))
blanchet@49331
    36
blanchet@49331
    37
fun has_thy thy th =
blanchet@49331
    38
  Context.theory_name thy = Context.theory_name (theory_of_thm th)
blanchet@49331
    39
blanchet@49546
    40
fun all_non_technical_facts ctxt thy =
blanchet@49546
    41
  let val css = Sledgehammer_Fact.clasimpset_rule_table_of ctxt in
blanchet@49546
    42
    all_facts (Proof_Context.init_global thy) false Symtab.empty [] [] css
blanchet@49546
    43
    |> filter_out (is_thm_technical o snd)
blanchet@49546
    44
  end
blanchet@49546
    45
blanchet@49331
    46
fun parent_facts thy thy_map =
blanchet@49331
    47
  let
blanchet@49331
    48
    fun add_last thy =
blanchet@49331
    49
      case AList.lookup (op =) thy_map (Context.theory_name thy) of
blanchet@49331
    50
        SOME (last_fact :: _) => insert (op =) last_fact
blanchet@49331
    51
      | _ => add_parent thy
blanchet@49331
    52
    and add_parent thy = fold add_last (Theory.parents_of thy)
blanchet@49331
    53
  in add_parent thy [] end
blanchet@49331
    54
blanchet@49331
    55
val thy_name_of_fact = hd o Long_Name.explode
blanchet@49319
    56
fun thy_of_fact thy = Context.this_theory thy o thy_name_of_fact
blanchet@49319
    57
blanchet@49453
    58
val all_names = map (rpair () o nickname_of) #> Symtab.make
blanchet@49331
    59
blanchet@49544
    60
fun smart_dependencies_of ctxt params prover facts all_names th =
blanchet@49544
    61
  case atp_dependencies_of ctxt params prover 0 facts all_names th of
blanchet@49544
    62
    SOME deps => deps
blanchet@49544
    63
  | NONE => isar_dependencies_of all_names th |> these
blanchet@49544
    64
blanchet@49545
    65
fun generate_accessibility ctxt thy include_thy file_name =
blanchet@49319
    66
  let
blanchet@49319
    67
    val path = file_name |> Path.explode
blanchet@49319
    68
    val _ = File.write path ""
blanchet@49319
    69
    fun do_fact fact prevs =
blanchet@49319
    70
      let
blanchet@49319
    71
        val s = escape_meta fact ^ ": " ^ escape_metas prevs ^ "\n"
blanchet@49319
    72
        val _ = File.append path s
blanchet@49319
    73
      in [fact] end
blanchet@49330
    74
    val thy_map =
blanchet@49546
    75
      all_non_technical_facts ctxt thy
blanchet@49319
    76
      |> not include_thy ? filter_out (has_thy thy o snd)
blanchet@49330
    77
      |> thy_map_from_facts
blanchet@49319
    78
    fun do_thy facts =
blanchet@49319
    79
      let
blanchet@49319
    80
        val thy = thy_of_fact thy (hd facts)
blanchet@49330
    81
        val parents = parent_facts thy thy_map
blanchet@49545
    82
      in fold_rev do_fact facts parents; () end
blanchet@49348
    83
  in fold_rev (fn (_, facts) => fn () => do_thy facts) thy_map () end
blanchet@49319
    84
blanchet@49333
    85
fun generate_features ctxt prover thy include_thy file_name =
blanchet@49319
    86
  let
blanchet@49319
    87
    val path = file_name |> Path.explode
blanchet@49319
    88
    val _ = File.write path ""
blanchet@49319
    89
    val facts =
blanchet@49546
    90
      all_non_technical_facts ctxt thy
blanchet@49319
    91
      |> not include_thy ? filter_out (has_thy thy o snd)
blanchet@49400
    92
    fun do_fact ((_, stature), th) =
blanchet@49319
    93
      let
blanchet@49393
    94
        val name = nickname_of th
blanchet@49333
    95
        val feats =
blanchet@49400
    96
          features_of ctxt prover (theory_of_thm th) stature [prop_of th]
blanchet@49319
    97
        val s = escape_meta name ^ ": " ^ escape_metas feats ^ "\n"
blanchet@49319
    98
      in File.append path s end
blanchet@49319
    99
  in List.app do_fact facts end
blanchet@49319
   100
blanchet@49545
   101
fun generate_isar_dependencies ctxt thy include_thy file_name =
blanchet@49319
   102
  let
blanchet@49319
   103
    val path = file_name |> Path.explode
blanchet@49319
   104
    val _ = File.write path ""
blanchet@49319
   105
    val ths =
blanchet@49546
   106
      all_non_technical_facts ctxt thy
blanchet@49319
   107
      |> not include_thy ? filter_out (has_thy thy o snd)
blanchet@49319
   108
      |> map snd
blanchet@49339
   109
    val all_names = all_names ths
blanchet@49319
   110
    fun do_thm th =
blanchet@49319
   111
      let
blanchet@49393
   112
        val name = nickname_of th
blanchet@49419
   113
        val deps = isar_dependencies_of all_names th |> these
blanchet@49319
   114
        val s = escape_meta name ^ ": " ^ escape_metas deps ^ "\n"
blanchet@49319
   115
      in File.append path s end
blanchet@49319
   116
  in List.app do_thm ths end
blanchet@49319
   117
blanchet@49407
   118
fun generate_atp_dependencies ctxt (params as {provers, ...}) thy include_thy
blanchet@49407
   119
                              file_name =
blanchet@49319
   120
  let
blanchet@49319
   121
    val path = file_name |> Path.explode
blanchet@49319
   122
    val _ = File.write path ""
blanchet@49333
   123
    val prover = hd provers
blanchet@49319
   124
    val facts =
blanchet@49546
   125
      all_non_technical_facts ctxt thy
blanchet@49319
   126
      |> not include_thy ? filter_out (has_thy thy o snd)
blanchet@49319
   127
    val ths = facts |> map snd
blanchet@49339
   128
    val all_names = all_names ths
blanchet@49319
   129
    fun do_thm th =
blanchet@49319
   130
      let
blanchet@49393
   131
        val name = nickname_of th
blanchet@49544
   132
        val deps = smart_dependencies_of ctxt params prover facts all_names th
blanchet@49319
   133
        val s = escape_meta name ^ ": " ^ escape_metas deps ^ "\n"
blanchet@49319
   134
      in File.append path s end
blanchet@49319
   135
  in List.app do_thm ths end
blanchet@49319
   136
blanchet@49544
   137
fun generate_commands ctxt (params as {provers, ...}) thy file_name =
blanchet@49249
   138
  let
blanchet@49249
   139
    val path = file_name |> Path.explode
blanchet@49249
   140
    val _ = File.write path ""
blanchet@49544
   141
    val prover = hd provers
blanchet@49546
   142
    val facts = all_non_technical_facts ctxt thy
blanchet@49249
   143
    val (new_facts, old_facts) =
blanchet@49249
   144
      facts |> List.partition (has_thy thy o snd)
blanchet@49249
   145
            |>> sort (thm_ord o pairself snd)
blanchet@49249
   146
    val ths = facts |> map snd
blanchet@49339
   147
    val all_names = all_names ths
blanchet@49400
   148
    fun do_fact ((_, stature), th) prevs =
blanchet@49249
   149
      let
blanchet@49393
   150
        val name = nickname_of th
blanchet@49400
   151
        val feats = features_of ctxt prover thy stature [prop_of th]
blanchet@49544
   152
        val deps = smart_dependencies_of ctxt params prover facts all_names th
blanchet@49249
   153
        val kind = Thm.legacy_get_kind th
blanchet@49544
   154
        val core =
blanchet@49544
   155
          escape_meta name ^ ": " ^ escape_metas prevs ^ "; " ^
blanchet@49544
   156
          escape_metas feats
blanchet@49249
   157
        val query = if kind <> "" then "? " ^ core ^ "\n" else ""
blanchet@49544
   158
        val update = "! " ^ core ^ "; " ^ escape_metas deps ^ "\n"
blanchet@49249
   159
        val _ = File.append path (query ^ update)
blanchet@49249
   160
      in [name] end
blanchet@49330
   161
    val thy_map = old_facts |> thy_map_from_facts
blanchet@49330
   162
    val parents = parent_facts thy thy_map
blanchet@49254
   163
  in fold do_fact new_facts parents; () end
blanchet@49254
   164
blanchet@49394
   165
fun generate_mepo_suggestions ctxt (params as {provers, ...}) thy max_relevant
blanchet@49307
   166
                              file_name =
blanchet@49254
   167
  let
blanchet@49254
   168
    val path = file_name |> Path.explode
blanchet@49254
   169
    val _ = File.write path ""
blanchet@49333
   170
    val prover = hd provers
blanchet@49546
   171
    val facts = all_non_technical_facts ctxt thy
blanchet@49254
   172
    val (new_facts, old_facts) =
blanchet@49254
   173
      facts |> List.partition (has_thy thy o snd)
blanchet@49254
   174
            |>> sort (thm_ord o pairself snd)
blanchet@49254
   175
    fun do_fact (fact as (_, th)) old_facts =
blanchet@49254
   176
      let
blanchet@49393
   177
        val name = nickname_of th
blanchet@49265
   178
        val goal = goal_of_thm thy th
blanchet@49307
   179
        val (_, hyp_ts, concl_t) = ATP_Util.strip_subgoal ctxt goal 1
blanchet@49254
   180
        val kind = Thm.legacy_get_kind th
blanchet@49254
   181
        val _ =
blanchet@49254
   182
          if kind <> "" then
blanchet@49254
   183
            let
blanchet@49254
   184
              val suggs =
blanchet@49307
   185
                old_facts
blanchet@49421
   186
                |> Sledgehammer_MePo.mepo_suggested_facts ctxt params prover
blanchet@49396
   187
                       max_relevant NONE hyp_ts concl_t
blanchet@49318
   188
                |> map (fn ((name, _), _) => name ())
blanchet@49318
   189
              val s = escape_meta name ^ ": " ^ escape_metas suggs ^ "\n"
blanchet@49254
   190
            in File.append path s end
blanchet@49254
   191
          else
blanchet@49254
   192
            ()
blanchet@49254
   193
      in fact :: old_facts end
blanchet@49254
   194
  in fold do_fact new_facts old_facts; () end
blanchet@49249
   195
blanchet@49249
   196
end;