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