src/HOL/TPTP/mash_export.ML
author blanchet
Sat, 15 Dec 2012 21:34:32 +0100
changeset 50559 89c0d2f13cca
parent 50523 0799339fea0f
child 50561 9a733bd6c0ba
permissions -rw-r--r--
MaSh exporter can now export subsets of the facts, as consecutive ranges
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 :
50559
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
    19
    Proof.context -> params -> int * int option -> theory list -> bool -> string
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
    20
    -> unit
50411
c9023d78d1a6 export ATP and Isar commands separately
blanchet
parents: 50402
diff changeset
    21
  val generate_isar_commands :
c9023d78d1a6 export ATP and Isar commands separately
blanchet
parents: 50402
diff changeset
    22
    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
    23
  val generate_prover_commands :
50559
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
    24
    Proof.context -> params -> int * int option -> theory list -> string -> unit
48379
2b5ad61e2ccc renamed "iter" fact filter to "MePo" (Meng--Paulson)
blanchet
parents: 48378
diff changeset
    25
  val generate_mepo_suggestions :
50349
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
    26
    Proof.context -> params -> theory list -> int -> string -> unit
48234
06216c789ac9 moved MaSh into own files
blanchet
parents:
diff changeset
    27
end;
06216c789ac9 moved MaSh into own files
blanchet
parents:
diff changeset
    28
06216c789ac9 moved MaSh into own files
blanchet
parents:
diff changeset
    29
structure MaSh_Export : MASH_EXPORT =
06216c789ac9 moved MaSh into own files
blanchet
parents:
diff changeset
    30
struct
06216c789ac9 moved MaSh into own files
blanchet
parents:
diff changeset
    31
50485
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50484
diff changeset
    32
open Sledgehammer_Fact
48381
1b7d798460bb renamed ML structures
blanchet
parents: 48379
diff changeset
    33
open Sledgehammer_MePo
1b7d798460bb renamed ML structures
blanchet
parents: 48379
diff changeset
    34
open Sledgehammer_MaSh
48245
854a47677335 generate ATP dependencies
blanchet
parents: 48242
diff changeset
    35
50559
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
    36
fun in_range (from, to) j =
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
    37
  j >= from andalso (to = NONE orelse j <= the to)
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
    38
48316
252f45c04042 drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents: 48315
diff changeset
    39
fun thy_map_from_facts ths =
50485
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50484
diff changeset
    40
  ths |> rev
48316
252f45c04042 drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents: 48315
diff changeset
    41
      |> 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
    42
      |> AList.coalesce (op =)
48378
9e96486d53ad handle local facts smoothly in MaSh
blanchet
parents: 48333
diff changeset
    43
      |> map (apsnd (map nickname_of))
48316
252f45c04042 drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents: 48315
diff changeset
    44
50349
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
    45
fun has_thm_thy th thy =
48316
252f45c04042 drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents: 48315
diff changeset
    46
  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
    47
50349
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
    48
fun has_thys thys th = exists (has_thm_thy th) thys
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
    49
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
    50
fun all_facts ctxt =
48531
7da5d3b8aef4 don't export technical theorems for MaSh
blanchet
parents: 48530
diff changeset
    51
  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
    52
    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
    53
    |> sort (thm_ord o pairself snd)
48531
7da5d3b8aef4 don't export technical theorems for MaSh
blanchet
parents: 48530
diff changeset
    54
  end
7da5d3b8aef4 don't export technical theorems for MaSh
blanchet
parents: 48530
diff changeset
    55
50349
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
    56
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
    57
  let
252f45c04042 drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents: 48315
diff changeset
    58
    fun add_last thy =
252f45c04042 drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents: 48315
diff changeset
    59
      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
    60
        SOME (last_fact :: _) => insert (op =) last_fact
252f45c04042 drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents: 48315
diff changeset
    61
      | _ => add_parent thy
252f45c04042 drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents: 48315
diff changeset
    62
    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
    63
  in add_parent thy end
48316
252f45c04042 drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents: 48315
diff changeset
    64
252f45c04042 drastic overhaul of MaSh data structures + fixed a few performance issues
blanchet
parents: 48315
diff changeset
    65
val thy_name_of_fact = hd o Long_Name.explode
50349
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
    66
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
    67
50349
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
    68
fun generate_accessibility ctxt thys include_thys file_name =
48304
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
    69
  let
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
    70
    val path = file_name |> Path.explode
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
    71
    val _ = File.write path ""
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
    72
    fun do_fact fact prevs =
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
    73
      let
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
    74
        val s = escape_meta fact ^ ": " ^ escape_metas prevs ^ "\n"
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
    75
        val _ = File.append path s
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
    76
      in [fact] end
48315
82d6e46c673f fixed order of accessibles + other tweaks to MaSh
blanchet
parents: 48304
diff changeset
    77
    val thy_map =
50349
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
    78
      all_facts ctxt
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
    79
      |> 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
    80
      |> thy_map_from_facts
48304
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
    81
    fun do_thy facts =
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
    82
      let
50349
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
    83
        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
    84
        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
    85
      in fold_rev do_fact facts parents; () end
48333
2250197977dc repair MaSh exporter
blanchet
parents: 48324
diff changeset
    86
  in fold_rev (fn (_, facts) => fn () => do_thy facts) thy_map () end
48304
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
    87
50349
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
    88
fun generate_features ctxt prover thys include_thys file_name =
48304
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
    89
  let
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
    90
    val path = file_name |> Path.explode
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
    91
    val _ = File.write path ""
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
    92
    val facts =
50349
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
    93
      all_facts ctxt
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
    94
      |> not include_thys ? filter_out (has_thys thys o snd)
48385
2779dea0b1e0 added locality as a MaSh feature
blanchet
parents: 48381
diff changeset
    95
    fun do_fact ((_, stature), th) =
48304
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
    96
      let
48378
9e96486d53ad handle local facts smoothly in MaSh
blanchet
parents: 48333
diff changeset
    97
        val name = nickname_of th
48318
325c8fd0d762 more consolidation of MaSh code
blanchet
parents: 48316
diff changeset
    98
        val feats =
48385
2779dea0b1e0 added locality as a MaSh feature
blanchet
parents: 48381
diff changeset
    99
          features_of ctxt prover (theory_of_thm th) stature [prop_of th]
50356
41f8f1f2e15d added feature weights in MaSh
blanchet
parents: 50349
diff changeset
   100
        val s = escape_meta name ^ ": " ^ encode_features feats ^ "\n"
48304
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
   101
      in File.append path s end
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
   102
  in List.app do_fact facts end
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
   103
50515
c4a27ab89c9b shared bad MaSh query detection between MePo and MaSh, so that the generated files mirror each other
blanchet
parents: 50511
diff changeset
   104
fun isar_or_prover_dependencies_of ctxt params_opt facts all_names th
c4a27ab89c9b shared bad MaSh query detection between MePo and MaSh, so that the generated files mirror each other
blanchet
parents: 50511
diff changeset
   105
                                   isar_deps_opt =
50411
c9023d78d1a6 export ATP and Isar commands separately
blanchet
parents: 50402
diff changeset
   106
  (case params_opt of
c9023d78d1a6 export ATP and Isar commands separately
blanchet
parents: 50402
diff changeset
   107
     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
   108
     prover_dependencies_of ctxt params prover 0 facts all_names th |> snd
50515
c4a27ab89c9b shared bad MaSh query detection between MePo and MaSh, so that the generated files mirror each other
blanchet
parents: 50511
diff changeset
   109
   | NONE =>
c4a27ab89c9b shared bad MaSh query detection between MePo and MaSh, so that the generated files mirror each other
blanchet
parents: 50511
diff changeset
   110
     case isar_deps_opt of
c4a27ab89c9b shared bad MaSh query detection between MePo and MaSh, so that the generated files mirror each other
blanchet
parents: 50511
diff changeset
   111
       SOME deps => deps
c4a27ab89c9b shared bad MaSh query detection between MePo and MaSh, so that the generated files mirror each other
blanchet
parents: 50511
diff changeset
   112
     | NONE => isar_dependencies_of all_names th)
50411
c9023d78d1a6 export ATP and Isar commands separately
blanchet
parents: 50402
diff changeset
   113
  |> these
c9023d78d1a6 export ATP and Isar commands separately
blanchet
parents: 50402
diff changeset
   114
50559
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
   115
fun generate_isar_or_prover_dependencies ctxt params_opt range thys include_thys
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
   116
                                         file_name =
48304
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
   117
  let
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
   118
    val path = file_name |> Path.explode
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
   119
    val facts =
50485
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50484
diff changeset
   120
      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
   121
    val all_names = build_all_names nickname_of facts
50559
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
   122
    fun do_fact (j, (_, th)) =
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
   123
      if in_range range j then
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
   124
        let
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
   125
          val name = nickname_of th
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
   126
          val deps =
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
   127
            isar_or_prover_dependencies_of ctxt params_opt facts all_names th
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
   128
                                           NONE
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
   129
        in escape_meta name ^ ": " ^ escape_metas deps ^ "\n" end
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
   130
      else
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
   131
        ""
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
   132
    val lines = Par_List.map do_fact (tag_list 1 facts)
50519
2951841ec011 parallelized MaSh exporter
blanchet
parents: 50515
diff changeset
   133
  in File.write_list path lines end
48304
50e64af9c829 more work on MaSh
blanchet
parents: 48303
diff changeset
   134
50411
c9023d78d1a6 export ATP and Isar commands separately
blanchet
parents: 50402
diff changeset
   135
fun generate_isar_dependencies ctxt =
50559
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
   136
  generate_isar_or_prover_dependencies ctxt NONE (1, NONE)
50411
c9023d78d1a6 export ATP and Isar commands separately
blanchet
parents: 50402
diff changeset
   137
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
   138
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
   139
  generate_isar_or_prover_dependencies ctxt (SOME params)
50411
c9023d78d1a6 export ATP and Isar commands separately
blanchet
parents: 50402
diff changeset
   140
50515
c4a27ab89c9b shared bad MaSh query detection between MePo and MaSh, so that the generated files mirror each other
blanchet
parents: 50511
diff changeset
   141
fun is_bad_query ctxt ho_atp th isar_deps =
c4a27ab89c9b shared bad MaSh query detection between MePo and MaSh, so that the generated files mirror each other
blanchet
parents: 50511
diff changeset
   142
  Thm.legacy_get_kind th = "" orelse null isar_deps orelse
50523
0799339fea0f get rid of some junk facts in the MaSh evaluation driver
blanchet
parents: 50519
diff changeset
   143
  is_blacklisted_or_something ctxt ho_atp (Thm.get_name_hint th)
50515
c4a27ab89c9b shared bad MaSh query detection between MePo and MaSh, so that the generated files mirror each other
blanchet
parents: 50511
diff changeset
   144
50559
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
   145
fun generate_isar_or_prover_commands ctxt prover params_opt range thys
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
   146
                                     file_name =
48234
06216c789ac9 moved MaSh into own files
blanchet
parents:
diff changeset
   147
  let
50511
8825c36cb1ce don't query blacklisted theorems in evaluation driver
blanchet
parents: 50485
diff changeset
   148
    val ho_atp = Sledgehammer_Provers.is_ho_atp ctxt prover
48234
06216c789ac9 moved MaSh into own files
blanchet
parents:
diff changeset
   149
    val path = file_name |> Path.explode
50349
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
   150
    val facts = all_facts ctxt
50485
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50484
diff changeset
   151
    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
   152
    val all_names = build_all_names nickname_of facts
50559
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
   153
    fun do_fact (j, ((name, ((_, stature), th)), prevs)) =
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
   154
      if in_range range j then
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
   155
        let
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
   156
          val feats =
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
   157
            features_of ctxt prover (theory_of_thm th) stature [prop_of th]
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
   158
          val isar_deps = isar_dependencies_of all_names th
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
   159
          val deps =
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
   160
            isar_or_prover_dependencies_of ctxt params_opt facts all_names th
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
   161
                                           (SOME isar_deps)
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
   162
          val core =
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
   163
            escape_meta name ^ ": " ^ escape_metas prevs ^ "; " ^
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
   164
            encode_features feats
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
   165
          val query =
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
   166
            if is_bad_query ctxt ho_atp th (these isar_deps) then ""
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
   167
            else "? " ^ core ^ "\n"
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
   168
          val update = "! " ^ core ^ "; " ^ escape_metas deps ^ "\n"
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
   169
        in query ^ update end
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
   170
      else
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
   171
        ""
48315
82d6e46c673f fixed order of accessibles + other tweaks to MaSh
blanchet
parents: 48304
diff changeset
   172
    val thy_map = old_facts |> thy_map_from_facts
50349
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
   173
    val parents = fold (add_thy_parent_facts thy_map) thys []
50519
2951841ec011 parallelized MaSh exporter
blanchet
parents: 50515
diff changeset
   174
    val new_facts = new_facts |> map (`(nickname_of o snd))
2951841ec011 parallelized MaSh exporter
blanchet
parents: 50515
diff changeset
   175
    val prevss = fst (split_last (parents :: map (single o fst) new_facts))
50559
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
   176
    val lines = Par_List.map do_fact (tag_list 1 (new_facts ~~ prevss))
50519
2951841ec011 parallelized MaSh exporter
blanchet
parents: 50515
diff changeset
   177
  in File.write_list path lines end
48239
0016290f904c generate Meng--Paulson facts for evaluation purposes
blanchet
parents: 48235
diff changeset
   178
50411
c9023d78d1a6 export ATP and Isar commands separately
blanchet
parents: 50402
diff changeset
   179
fun generate_isar_commands ctxt prover =
50559
89c0d2f13cca MaSh exporter can now export subsets of the facts, as consecutive ranges
blanchet
parents: 50523
diff changeset
   180
  generate_isar_or_prover_commands ctxt prover NONE (1, NONE)
50411
c9023d78d1a6 export ATP and Isar commands separately
blanchet
parents: 50402
diff changeset
   181
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
   182
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
   183
  generate_isar_or_prover_commands ctxt prover (SOME params)
50411
c9023d78d1a6 export ATP and Isar commands separately
blanchet
parents: 50402
diff changeset
   184
50515
c4a27ab89c9b shared bad MaSh query detection between MePo and MaSh, so that the generated files mirror each other
blanchet
parents: 50511
diff changeset
   185
fun generate_mepo_suggestions ctxt (params as {provers = prover :: _, ...}) thys
c4a27ab89c9b shared bad MaSh query detection between MePo and MaSh, so that the generated files mirror each other
blanchet
parents: 50511
diff changeset
   186
                              max_facts file_name =
48239
0016290f904c generate Meng--Paulson facts for evaluation purposes
blanchet
parents: 48235
diff changeset
   187
  let
50515
c4a27ab89c9b shared bad MaSh query detection between MePo and MaSh, so that the generated files mirror each other
blanchet
parents: 50511
diff changeset
   188
    val ho_atp = Sledgehammer_Provers.is_ho_atp ctxt prover
48239
0016290f904c generate Meng--Paulson facts for evaluation purposes
blanchet
parents: 48235
diff changeset
   189
    val path = file_name |> Path.explode
50349
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
   190
    val facts = all_facts ctxt
50485
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50484
diff changeset
   191
    val (new_facts, old_facts) = facts |> List.partition (has_thys thys o snd)
50515
c4a27ab89c9b shared bad MaSh query detection between MePo and MaSh, so that the generated files mirror each other
blanchet
parents: 50511
diff changeset
   192
    val all_names = build_all_names nickname_of facts
50523
0799339fea0f get rid of some junk facts in the MaSh evaluation driver
blanchet
parents: 50519
diff changeset
   193
    fun do_fact ((_, th), old_facts) =
48239
0016290f904c generate Meng--Paulson facts for evaluation purposes
blanchet
parents: 48235
diff changeset
   194
      let
48378
9e96486d53ad handle local facts smoothly in MaSh
blanchet
parents: 48333
diff changeset
   195
        val name = nickname_of th
50349
b79803ee14f3 generalized MaSh exporter to sets of theories
blanchet
parents: 48667
diff changeset
   196
        val goal = goal_of_thm (Proof_Context.theory_of ctxt) th
48292
7fcee834c7f5 more code rationalization in relevance filter
blanchet
parents: 48289
diff changeset
   197
        val (_, hyp_ts, concl_t) = ATP_Util.strip_subgoal ctxt goal 1
50515
c4a27ab89c9b shared bad MaSh query detection between MePo and MaSh, so that the generated files mirror each other
blanchet
parents: 50511
diff changeset
   198
        val isar_deps = isar_dependencies_of all_names th
50519
2951841ec011 parallelized MaSh exporter
blanchet
parents: 50515
diff changeset
   199
      in
2951841ec011 parallelized MaSh exporter
blanchet
parents: 50515
diff changeset
   200
        if is_bad_query ctxt ho_atp th (these isar_deps) then
2951841ec011 parallelized MaSh exporter
blanchet
parents: 50515
diff changeset
   201
          ""
2951841ec011 parallelized MaSh exporter
blanchet
parents: 50515
diff changeset
   202
        else
2951841ec011 parallelized MaSh exporter
blanchet
parents: 50515
diff changeset
   203
          let
2951841ec011 parallelized MaSh exporter
blanchet
parents: 50515
diff changeset
   204
            val suggs =
2951841ec011 parallelized MaSh exporter
blanchet
parents: 50515
diff changeset
   205
              old_facts
2951841ec011 parallelized MaSh exporter
blanchet
parents: 50515
diff changeset
   206
              |> Sledgehammer_MePo.mepo_suggested_facts ctxt params prover
2951841ec011 parallelized MaSh exporter
blanchet
parents: 50515
diff changeset
   207
                     max_facts NONE hyp_ts concl_t
2951841ec011 parallelized MaSh exporter
blanchet
parents: 50515
diff changeset
   208
              |> map (nickname_of o snd)
2951841ec011 parallelized MaSh exporter
blanchet
parents: 50515
diff changeset
   209
          in escape_meta name ^ ": " ^ escape_metas suggs ^ "\n" end
2951841ec011 parallelized MaSh exporter
blanchet
parents: 50515
diff changeset
   210
      end
2951841ec011 parallelized MaSh exporter
blanchet
parents: 50515
diff changeset
   211
    fun accum x (yss as ys :: _) = (x :: ys) :: yss
2951841ec011 parallelized MaSh exporter
blanchet
parents: 50515
diff changeset
   212
    val old_factss = tl (fold accum new_facts [old_facts])
2951841ec011 parallelized MaSh exporter
blanchet
parents: 50515
diff changeset
   213
    val lines = Par_List.map do_fact (new_facts ~~ rev old_factss)
2951841ec011 parallelized MaSh exporter
blanchet
parents: 50515
diff changeset
   214
  in File.write_list path lines end
48234
06216c789ac9 moved MaSh into own files
blanchet
parents:
diff changeset
   215
06216c789ac9 moved MaSh into own files
blanchet
parents:
diff changeset
   216
end;