src/HOL/TPTP/mash_export.ML
author blanchet
Wed, 12 Dec 2012 21:48:29 +0100
changeset 50511 8825c36cb1ce
parent 50485 3c6ac2da2f45
child 50515 c4a27ab89c9b
permissions -rw-r--r--
don't query blacklisted theorems in evaluation driver
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
48234
06216c789ac9 moved MaSh into own files
blanchet
parents:
diff changeset
     1
(*  Title:      HOL/TPTP/mash_export.ML
06216c789ac9 moved MaSh into own files
blanchet
parents:
diff changeset
     2
    Author:     Jasmin Blanchette, TU Muenchen
06216c789ac9 moved MaSh into own files
blanchet
parents:
diff changeset
     3
    Copyright   2012
06216c789ac9 moved MaSh into own files
blanchet
parents:
diff changeset
     4
06216c789ac9 moved MaSh into own files
blanchet
parents:
diff changeset
     5
Export Isabelle theory information for MaSh (Machine-learning for Sledgehammer).
06216c789ac9 moved MaSh into own files
blanchet
parents:
diff changeset
     6
*)
06216c789ac9 moved MaSh into own files
blanchet
parents:
diff changeset
     7
06216c789ac9 moved MaSh into own files
blanchet
parents:
diff changeset
     8
signature MASH_EXPORT =
06216c789ac9 moved MaSh into own files
blanchet
parents:
diff changeset
     9
sig
48251
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 48250
diff changeset
    10
  type params = Sledgehammer_Provers.params
48235
40655464a93b MaSh evaluation driver
blanchet
parents: 48234
diff changeset
    11
50349
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
    12
  val generate_accessibility :
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
    13
    Proof.context -> theory list -> bool -> string -> unit
48318
325c8fd0d762 more consolidation of MaSh code
blanchet
parents: 48316
diff changeset
    14
  val generate_features :
50349
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
    15
    Proof.context -> string -> theory list -> bool -> string -> unit
48333
2250197977dc repair MaSh exporter
blanchet
parents: 48324
diff changeset
    16
  val generate_isar_dependencies :
50349
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
    17
    Proof.context -> theory list -> bool -> string -> unit
50484
8ec31bdb9d36 adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
blanchet
parents: 50442
diff changeset
    18
  val generate_prover_dependencies :
50349
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
    19
    Proof.context -> params -> theory list -> bool -> string -> unit
50411
c9023d78d1a6 export ATP and Isar commands separately
blanchet
parents: 50402
diff changeset
    20
  val generate_isar_commands :
c9023d78d1a6 export ATP and Isar commands separately
blanchet
parents: 50402
diff changeset
    21
    Proof.context -> string -> theory list -> string -> unit
50484
8ec31bdb9d36 adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
blanchet
parents: 50442
diff changeset
    22
  val generate_prover_commands :
50349
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
    23
    Proof.context -> params -> theory list -> string -> unit
48379
2b5ad61e2ccc renamed "iter" fact filter to "MePo" (Meng--Paulson)
blanchet
parents: 48378
diff changeset
    24
  val generate_mepo_suggestions :
50349
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
    25
    Proof.context -> params -> theory list -> int -> string -> unit
48234
06216c789ac9 moved MaSh into own files
blanchet
parents:
diff changeset
    26
end;
06216c789ac9 moved MaSh into own files
blanchet
parents:
diff changeset
    27
06216c789ac9 moved MaSh into own files
blanchet
parents:
diff changeset
    28
structure MaSh_Export : MASH_EXPORT =
06216c789ac9 moved MaSh into own files
blanchet
parents:
diff changeset
    29
struct
06216c789ac9 moved MaSh into own files
blanchet
parents:
diff changeset
    30
50485
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50484
diff changeset
    31
open Sledgehammer_Fact
48381
1b7d798460bb renamed ML structures
blanchet
parents: 48379
diff changeset
    32
open Sledgehammer_MePo
1b7d798460bb renamed ML structures
blanchet
parents: 48379
diff changeset
    33
open Sledgehammer_MaSh
48245
854a47677335 generate ATP dependencies
blanchet
parents: 48242
diff changeset
    34
48316
252f45c04042 drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents: 48315
diff changeset
    35
fun thy_map_from_facts ths =
50485
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50484
diff changeset
    36
  ths |> rev
48316
252f45c04042 drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents: 48315
diff changeset
    37
      |> map (snd #> `(theory_of_thm #> Context.theory_name))
252f45c04042 drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents: 48315
diff changeset
    38
      |> AList.coalesce (op =)
48378
9e96486d53ad handle local facts smoothly in MaSh
blanchet
parents: 48333
diff changeset
    39
      |> map (apsnd (map nickname_of))
48316
252f45c04042 drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents: 48315
diff changeset
    40
50349
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
    41
fun has_thm_thy th thy =
48316
252f45c04042 drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents: 48315
diff changeset
    42
  Context.theory_name thy = Context.theory_name (theory_of_thm th)
252f45c04042 drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents: 48315
diff changeset
    43
50349
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
    44
fun has_thys thys th = exists (has_thm_thy th) thys
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
    45
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
    46
fun all_facts ctxt =
48531
7da5d3b8aef4 don't export technical theorems for MaSh
blanchet
parents: 48530
diff changeset
    47
  let val css = Sledgehammer_Fact.clasimpset_rule_table_of ctxt in
50442
4f6a4d32522c don't blacklist "case" theorems -- this causes problems in MaSh later
blanchet
parents: 50434
diff changeset
    48
    Sledgehammer_Fact.all_facts ctxt true false Symtab.empty [] [] css
50485
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50484
diff changeset
    49
    |> sort (thm_ord o pairself snd)
48531
7da5d3b8aef4 don't export technical theorems for MaSh
blanchet
parents: 48530
diff changeset
    50
  end
7da5d3b8aef4 don't export technical theorems for MaSh
blanchet
parents: 48530
diff changeset
    51
50349
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
    52
fun add_thy_parent_facts thy_map thy =
48316
252f45c04042 drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents: 48315
diff changeset
    53
  let
252f45c04042 drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents: 48315
diff changeset
    54
    fun add_last thy =
252f45c04042 drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents: 48315
diff changeset
    55
      case AList.lookup (op =) thy_map (Context.theory_name thy) of
252f45c04042 drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents: 48315
diff changeset
    56
        SOME (last_fact :: _) => insert (op =) last_fact
252f45c04042 drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents: 48315
diff changeset
    57
      | _ => add_parent thy
252f45c04042 drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents: 48315
diff changeset
    58
    and add_parent thy = fold add_last (Theory.parents_of thy)
50349
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
    59
  in add_parent thy end
48316
252f45c04042 drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents: 48315
diff changeset
    60
252f45c04042 drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents: 48315
diff changeset
    61
val thy_name_of_fact = hd o Long_Name.explode
50349
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
    62
fun thy_of_fact thy = Context.get_theory thy o thy_name_of_fact
48304
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
    63
50349
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
    64
fun generate_accessibility ctxt thys include_thys file_name =
48304
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
    65
  let
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
    66
    val path = file_name |> Path.explode
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
    67
    val _ = File.write path ""
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
    68
    fun do_fact fact prevs =
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
    69
      let
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
    70
        val s = escape_meta fact ^ ": " ^ escape_metas prevs ^ "\n"
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
    71
        val _ = File.append path s
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
    72
      in [fact] end
48315
82d6e46c673f fixed order of accessibles + other tweaks to MaSh
blanchet
parents: 48304
diff changeset
    73
    val thy_map =
50349
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
    74
      all_facts ctxt
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
    75
      |> not include_thys ? filter_out (has_thys thys o snd)
48315
82d6e46c673f fixed order of accessibles + other tweaks to MaSh
blanchet
parents: 48304
diff changeset
    76
      |> thy_map_from_facts
48304
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
    77
    fun do_thy facts =
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
    78
      let
50349
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
    79
        val thy = thy_of_fact (Proof_Context.theory_of ctxt) (hd facts)
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
    80
        val parents = add_thy_parent_facts thy_map thy []
48530
d443166f9520 repaired accessibility chains generated by MaSh exporter + tuned one function out
blanchet
parents: 48529
diff changeset
    81
      in fold_rev do_fact facts parents; () end
48333
2250197977dc repair MaSh exporter
blanchet
parents: 48324
diff changeset
    82
  in fold_rev (fn (_, facts) => fn () => do_thy facts) thy_map () end
48304
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
    83
50349
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
    84
fun generate_features ctxt prover thys include_thys file_name =
48304
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
    85
  let
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
    86
    val path = file_name |> Path.explode
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
    87
    val _ = File.write path ""
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
    88
    val facts =
50349
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
    89
      all_facts ctxt
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
    90
      |> not include_thys ? filter_out (has_thys thys o snd)
48385
2779dea0b1e0 added locality as a MaSh feature
blanchet
parents: 48381
diff changeset
    91
    fun do_fact ((_, stature), th) =
48304
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
    92
      let
48378
9e96486d53ad handle local facts smoothly in MaSh
blanchet
parents: 48333
diff changeset
    93
        val name = nickname_of th
48318
325c8fd0d762 more consolidation of MaSh code
blanchet
parents: 48316
diff changeset
    94
        val feats =
48385
2779dea0b1e0 added locality as a MaSh feature
blanchet
parents: 48381
diff changeset
    95
          features_of ctxt prover (theory_of_thm th) stature [prop_of th]
50356
41f8f1f2e15d added feature weights in MaSh
blanchet
parents: 50349
diff changeset
    96
        val s = escape_meta name ^ ": " ^ encode_features feats ^ "\n"
48304
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
    97
      in File.append path s end
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
    98
  in List.app do_fact facts end
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
    99
50484
8ec31bdb9d36 adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
blanchet
parents: 50442
diff changeset
   100
fun isar_or_prover_dependencies_of ctxt params_opt facts all_names th =
50411
c9023d78d1a6 export ATP and Isar commands separately
blanchet
parents: 50402
diff changeset
   101
  (case params_opt of
c9023d78d1a6 export ATP and Isar commands separately
blanchet
parents: 50402
diff changeset
   102
     SOME (params as {provers = prover :: _, ...}) =>
50484
8ec31bdb9d36 adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
blanchet
parents: 50442
diff changeset
   103
     prover_dependencies_of ctxt params prover 0 facts all_names th |> snd
50411
c9023d78d1a6 export ATP and Isar commands separately
blanchet
parents: 50402
diff changeset
   104
   | NONE => isar_dependencies_of all_names th)
c9023d78d1a6 export ATP and Isar commands separately
blanchet
parents: 50402
diff changeset
   105
  |> these
c9023d78d1a6 export ATP and Isar commands separately
blanchet
parents: 50402
diff changeset
   106
50484
8ec31bdb9d36 adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
blanchet
parents: 50442
diff changeset
   107
fun generate_isar_or_prover_dependencies ctxt params_opt thys include_thys
8ec31bdb9d36 adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
blanchet
parents: 50442
diff changeset
   108
                                         file_name =
48304
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
   109
  let
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
   110
    val path = file_name |> Path.explode
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
   111
    val _ = File.write path ""
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
   112
    val facts =
50485
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50484
diff changeset
   113
      all_facts ctxt |> not include_thys ? filter_out (has_thys thys o snd)
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50484
diff changeset
   114
    val all_names = build_all_names nickname_of facts
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50484
diff changeset
   115
    fun do_fact (_, th) =
48304
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
   116
      let
48378
9e96486d53ad handle local facts smoothly in MaSh
blanchet
parents: 48333
diff changeset
   117
        val name = nickname_of th
48665
14b0732c72f7 crank up max number of dependencies
blanchet
parents: 48531
diff changeset
   118
        val deps =
50484
8ec31bdb9d36 adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
blanchet
parents: 50442
diff changeset
   119
          isar_or_prover_dependencies_of ctxt params_opt facts all_names th
48304
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
   120
        val s = escape_meta name ^ ": " ^ escape_metas deps ^ "\n"
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
   121
      in File.append path s end
50485
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50484
diff changeset
   122
  in List.app do_fact facts end
48304
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
   123
50411
c9023d78d1a6 export ATP and Isar commands separately
blanchet
parents: 50402
diff changeset
   124
fun generate_isar_dependencies ctxt =
50484
8ec31bdb9d36 adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
blanchet
parents: 50442
diff changeset
   125
  generate_isar_or_prover_dependencies ctxt NONE
50411
c9023d78d1a6 export ATP and Isar commands separately
blanchet
parents: 50402
diff changeset
   126
50484
8ec31bdb9d36 adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
blanchet
parents: 50442
diff changeset
   127
fun generate_prover_dependencies ctxt params =
8ec31bdb9d36 adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
blanchet
parents: 50442
diff changeset
   128
  generate_isar_or_prover_dependencies ctxt (SOME params)
50411
c9023d78d1a6 export ATP and Isar commands separately
blanchet
parents: 50402
diff changeset
   129
50484
8ec31bdb9d36 adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
blanchet
parents: 50442
diff changeset
   130
fun generate_isar_or_prover_commands ctxt prover params_opt thys file_name =
48234
06216c789ac9 moved MaSh into own files
blanchet
parents:
diff changeset
   131
  let
50511
8825c36cb1ce don't query blacklisted theorems in evaluation driver
blanchet
parents: 50485
diff changeset
   132
    val ho_atp = Sledgehammer_Provers.is_ho_atp ctxt prover
48234
06216c789ac9 moved MaSh into own files
blanchet
parents:
diff changeset
   133
    val path = file_name |> Path.explode
06216c789ac9 moved MaSh into own files
blanchet
parents:
diff changeset
   134
    val _ = File.write path ""
50349
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
   135
    val facts = all_facts ctxt
50485
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50484
diff changeset
   136
    val (new_facts, old_facts) = facts |> List.partition (has_thys thys o snd)
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50484
diff changeset
   137
    val all_names = build_all_names nickname_of facts
48385
2779dea0b1e0 added locality as a MaSh feature
blanchet
parents: 48381
diff changeset
   138
    fun do_fact ((_, stature), th) prevs =
48234
06216c789ac9 moved MaSh into own files
blanchet
parents:
diff changeset
   139
      let
48378
9e96486d53ad handle local facts smoothly in MaSh
blanchet
parents: 48333
diff changeset
   140
        val name = nickname_of th
50349
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
   141
        val feats =
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
   142
          features_of ctxt prover (theory_of_thm th) stature [prop_of th]
48665
14b0732c72f7 crank up max number of dependencies
blanchet
parents: 48531
diff changeset
   143
        val deps =
50484
8ec31bdb9d36 adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
blanchet
parents: 50442
diff changeset
   144
          isar_or_prover_dependencies_of ctxt params_opt facts all_names th
48234
06216c789ac9 moved MaSh into own files
blanchet
parents:
diff changeset
   145
        val kind = Thm.legacy_get_kind th
48529
716ec3458b1d generate fact name in queries again + use ATP dependencies when possible
blanchet
parents: 48438
diff changeset
   146
        val core =
716ec3458b1d generate fact name in queries again + use ATP dependencies when possible
blanchet
parents: 48438
diff changeset
   147
          escape_meta name ^ ": " ^ escape_metas prevs ^ "; " ^
50356
41f8f1f2e15d added feature weights in MaSh
blanchet
parents: 50349
diff changeset
   148
          encode_features feats
48666
0ba2e9be9f20 don't generate queries for empty dependencies
blanchet
parents: 48665
diff changeset
   149
        val query =
50511
8825c36cb1ce don't query blacklisted theorems in evaluation driver
blanchet
parents: 50485
diff changeset
   150
          if kind = "" orelse null deps orelse
8825c36cb1ce don't query blacklisted theorems in evaluation driver
blanchet
parents: 50485
diff changeset
   151
             is_blacklisted_or_something ctxt ho_atp name orelse
8825c36cb1ce don't query blacklisted theorems in evaluation driver
blanchet
parents: 50485
diff changeset
   152
             Sledgehammer_Fact.is_bad_for_atps ho_atp th then
8825c36cb1ce don't query blacklisted theorems in evaluation driver
blanchet
parents: 50485
diff changeset
   153
            ""
8825c36cb1ce don't query blacklisted theorems in evaluation driver
blanchet
parents: 50485
diff changeset
   154
          else
8825c36cb1ce don't query blacklisted theorems in evaluation driver
blanchet
parents: 50485
diff changeset
   155
            "? " ^ core ^ "\n"
48529
716ec3458b1d generate fact name in queries again + use ATP dependencies when possible
blanchet
parents: 48438
diff changeset
   156
        val update = "! " ^ core ^ "; " ^ escape_metas deps ^ "\n"
48234
06216c789ac9 moved MaSh into own files
blanchet
parents:
diff changeset
   157
        val _ = File.append path (query ^ update)
06216c789ac9 moved MaSh into own files
blanchet
parents:
diff changeset
   158
      in [name] end
48315
82d6e46c673f fixed order of accessibles + other tweaks to MaSh
blanchet
parents: 48304
diff changeset
   159
    val thy_map = old_facts |> thy_map_from_facts
50349
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
   160
    val parents = fold (add_thy_parent_facts thy_map) thys []
48239
0016290f904c generate Meng--Paulson facts for evaluation purposes
blanchet
parents: 48235
diff changeset
   161
  in fold do_fact new_facts parents; () end
0016290f904c generate Meng--Paulson facts for evaluation purposes
blanchet
parents: 48235
diff changeset
   162
50411
c9023d78d1a6 export ATP and Isar commands separately
blanchet
parents: 50402
diff changeset
   163
fun generate_isar_commands ctxt prover =
50484
8ec31bdb9d36 adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
blanchet
parents: 50442
diff changeset
   164
  generate_isar_or_prover_commands ctxt prover NONE
50411
c9023d78d1a6 export ATP and Isar commands separately
blanchet
parents: 50402
diff changeset
   165
50484
8ec31bdb9d36 adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
blanchet
parents: 50442
diff changeset
   166
fun generate_prover_commands ctxt (params as {provers = prover :: _, ...}) =
8ec31bdb9d36 adopt the neutral "prover" terminology for MaSh rather than the ambiguous/wrong ATP terminology (which sometimes excludes SMT solvers)
blanchet
parents: 50442
diff changeset
   167
  generate_isar_or_prover_commands ctxt prover (SOME params)
50411
c9023d78d1a6 export ATP and Isar commands separately
blanchet
parents: 50402
diff changeset
   168
50349
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
   169
fun generate_mepo_suggestions ctxt (params as {provers, ...}) thys max_relevant
48292
7fcee834c7f5 more code rationalization in relevance filter
blanchet
parents: 48289
diff changeset
   170
                              file_name =
48239
0016290f904c generate Meng--Paulson facts for evaluation purposes
blanchet
parents: 48235
diff changeset
   171
  let
0016290f904c generate Meng--Paulson facts for evaluation purposes
blanchet
parents: 48235
diff changeset
   172
    val path = file_name |> Path.explode
0016290f904c generate Meng--Paulson facts for evaluation purposes
blanchet
parents: 48235
diff changeset
   173
    val _ = File.write path ""
48318
325c8fd0d762 more consolidation of MaSh code
blanchet
parents: 48316
diff changeset
   174
    val prover = hd provers
50349
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
   175
    val facts = all_facts ctxt
50485
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50484
diff changeset
   176
    val (new_facts, old_facts) = facts |> List.partition (has_thys thys o snd)
48239
0016290f904c generate Meng--Paulson facts for evaluation purposes
blanchet
parents: 48235
diff changeset
   177
    fun do_fact (fact as (_, th)) old_facts =
0016290f904c generate Meng--Paulson facts for evaluation purposes
blanchet
parents: 48235
diff changeset
   178
      let
48378
9e96486d53ad handle local facts smoothly in MaSh
blanchet
parents: 48333
diff changeset
   179
        val name = nickname_of th
50349
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
   180
        val goal = goal_of_thm (Proof_Context.theory_of ctxt) th
48292
7fcee834c7f5 more code rationalization in relevance filter
blanchet
parents: 48289
diff changeset
   181
        val (_, hyp_ts, concl_t) = ATP_Util.strip_subgoal ctxt goal 1
48239
0016290f904c generate Meng--Paulson facts for evaluation purposes
blanchet
parents: 48235
diff changeset
   182
        val kind = Thm.legacy_get_kind th
0016290f904c generate Meng--Paulson facts for evaluation purposes
blanchet
parents: 48235
diff changeset
   183
        val _ =
0016290f904c generate Meng--Paulson facts for evaluation purposes
blanchet
parents: 48235
diff changeset
   184
          if kind <> "" then
0016290f904c generate Meng--Paulson facts for evaluation purposes
blanchet
parents: 48235
diff changeset
   185
            let
0016290f904c generate Meng--Paulson facts for evaluation purposes
blanchet
parents: 48235
diff changeset
   186
              val suggs =
48292
7fcee834c7f5 more code rationalization in relevance filter
blanchet
parents: 48289
diff changeset
   187
                old_facts
48406
b002cc16aa99 honor suggested MaSh weights
blanchet
parents: 48404
diff changeset
   188
                |> Sledgehammer_MePo.mepo_suggested_facts ctxt params prover
48381
1b7d798460bb renamed ML structures
blanchet
parents: 48379
diff changeset
   189
                       max_relevant NONE hyp_ts concl_t
50402
923f1e199f4f use right names in MePo exporter
blanchet
parents: 50356
diff changeset
   190
                |> map (nickname_of o snd)
48303
f1d135d0ea69 improved MaSh string escaping and make more operations string-based
blanchet
parents: 48302
diff changeset
   191
              val s = escape_meta name ^ ": " ^ escape_metas suggs ^ "\n"
48239
0016290f904c generate Meng--Paulson facts for evaluation purposes
blanchet
parents: 48235
diff changeset
   192
            in File.append path s end
0016290f904c generate Meng--Paulson facts for evaluation purposes
blanchet
parents: 48235
diff changeset
   193
          else
0016290f904c generate Meng--Paulson facts for evaluation purposes
blanchet
parents: 48235
diff changeset
   194
            ()
0016290f904c generate Meng--Paulson facts for evaluation purposes
blanchet
parents: 48235
diff changeset
   195
      in fact :: old_facts end
0016290f904c generate Meng--Paulson facts for evaluation purposes
blanchet
parents: 48235
diff changeset
   196
  in fold do_fact new_facts old_facts; () end
48234
06216c789ac9 moved MaSh into own files
blanchet
parents:
diff changeset
   197
06216c789ac9 moved MaSh into own files
blanchet
parents:
diff changeset
   198
end;