src/HOL/Tools/Sledgehammer/sledgehammer_fact.ML
author blanchet
Tue, 10 Sep 2013 15:56:51 +0200
changeset 53508 8d8f72aa5c0b
parent 53507 a6ed27399ba9
child 53509 cf7679195169
permissions -rw-r--r--
got rid of another taboo that appears to make no difference in practice (and that slows down the relevance filter)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
     1
(*  Title:      HOL/Tools/Sledgehammer/sledgehammer_fact.ML
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
     2
    Author:     Jia Meng, Cambridge University Computer Laboratory and NICTA
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
     3
    Author:     Jasmin Blanchette, TU Muenchen
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
     4
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
     5
Sledgehammer fact handling.
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
     6
*)
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
     7
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
     8
signature SLEDGEHAMMER_FACT =
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
     9
sig
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    10
  type status = ATP_Problem_Generate.status
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    11
  type stature = ATP_Problem_Generate.stature
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    12
51004
5f2788c38127 distinguish raw and non-raw facts, using raw for 10 000s of facts and non-raw after selection of some hundreds
blanchet
parents: 50815
diff changeset
    13
  type raw_fact = ((unit -> string) * stature) * thm
5f2788c38127 distinguish raw and non-raw facts, using raw for 10 000s of facts and non-raw after selection of some hundreds
blanchet
parents: 50815
diff changeset
    14
  type fact = (string * stature) * thm
48296
e7f01b7e244e gracefully handle the case of empty theories when going up the accessibility chain
blanchet
parents: 48292
diff changeset
    15
48292
7fcee834c7f5 more code rationalization in relevance filter
blanchet
parents: 48251
diff changeset
    16
  type fact_override =
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    17
    {add : (Facts.ref * Attrib.src list) list,
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    18
     del : (Facts.ref * Attrib.src list) list,
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    19
     only : bool}
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    20
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    21
  val ignore_no_atp : bool Config.T
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    22
  val instantiate_inducts : bool Config.T
48292
7fcee834c7f5 more code rationalization in relevance filter
blanchet
parents: 48251
diff changeset
    23
  val no_fact_override : fact_override
52031
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51998
diff changeset
    24
  val fact_of_ref :
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    25
    Proof.context -> unit Symtab.table -> thm list -> status Termtab.table
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    26
    -> Facts.ref * Attrib.src list -> ((string * stature) * thm) list
50047
45684acf0b94 thread context correctly when printing backquoted facts
blanchet
parents: 49981
diff changeset
    27
  val backquote_thm : Proof.context -> thm -> string
50511
8825c36cb1ce don't query blacklisted theorems in evaluation driver
blanchet
parents: 50510
diff changeset
    28
  val is_blacklisted_or_something : Proof.context -> bool -> string -> bool
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    29
  val clasimpset_rule_table_of : Proof.context -> status Termtab.table
50735
6b232d76cbc9 refined class handling, to prevent cycles in fact graph
blanchet
parents: 50734
diff changeset
    30
  val build_name_tables :
6b232d76cbc9 refined class handling, to prevent cycles in fact graph
blanchet
parents: 50734
diff changeset
    31
    (thm -> string) -> ('a * thm) list
6b232d76cbc9 refined class handling, to prevent cycles in fact graph
blanchet
parents: 50734
diff changeset
    32
    -> string Symtab.table * string Symtab.table
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    33
  val maybe_instantiate_inducts :
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    34
    Proof.context -> term list -> term -> (((unit -> string) * 'a) * thm) list
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    35
    -> (((unit -> string) * 'a) * thm) list
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    36
  val maybe_filter_no_atps : Proof.context -> ('a * thm) list -> ('a * thm) list
51004
5f2788c38127 distinguish raw and non-raw facts, using raw for 10 000s of facts and non-raw after selection of some hundreds
blanchet
parents: 50815
diff changeset
    37
  val fact_of_raw_fact : raw_fact -> fact
48530
d443166f9520 repaired accessibility chains generated by MaSh exporter + tuned one function out
blanchet
parents: 48440
diff changeset
    38
  val all_facts :
50442
4f6a4d32522c don't blacklist "case" theorems -- this causes problems in MaSh later
blanchet
parents: 50239
diff changeset
    39
    Proof.context -> bool -> bool -> unit Symtab.table -> thm list -> thm list
51004
5f2788c38127 distinguish raw and non-raw facts, using raw for 10 000s of facts and non-raw after selection of some hundreds
blanchet
parents: 50815
diff changeset
    40
    -> status Termtab.table -> raw_fact list
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    41
  val nearly_all_facts :
48299
5e5c6616f0fe centrally construct expensive data structures
blanchet
parents: 48298
diff changeset
    42
    Proof.context -> bool -> fact_override -> unit Symtab.table
51004
5f2788c38127 distinguish raw and non-raw facts, using raw for 10 000s of facts and non-raw after selection of some hundreds
blanchet
parents: 50815
diff changeset
    43
    -> status Termtab.table -> thm list -> term list -> term -> raw_fact list
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    44
end;
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    45
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    46
structure Sledgehammer_Fact : SLEDGEHAMMER_FACT =
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    47
struct
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    48
50495
bd9a0028b063 better tautology check -- don't reject "prod_cases3" for example
blanchet
parents: 50490
diff changeset
    49
open ATP_Util
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    50
open ATP_Problem_Generate
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    51
open Metis_Tactic
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    52
open Sledgehammer_Util
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    53
51004
5f2788c38127 distinguish raw and non-raw facts, using raw for 10 000s of facts and non-raw after selection of some hundreds
blanchet
parents: 50815
diff changeset
    54
type raw_fact = ((unit -> string) * stature) * thm
5f2788c38127 distinguish raw and non-raw facts, using raw for 10 000s of facts and non-raw after selection of some hundreds
blanchet
parents: 50815
diff changeset
    55
type fact = (string * stature) * thm
48296
e7f01b7e244e gracefully handle the case of empty theories when going up the accessibility chain
blanchet
parents: 48292
diff changeset
    56
48292
7fcee834c7f5 more code rationalization in relevance filter
blanchet
parents: 48251
diff changeset
    57
type fact_override =
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    58
  {add : (Facts.ref * Attrib.src list) list,
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    59
   del : (Facts.ref * Attrib.src list) list,
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    60
   only : bool}
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    61
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    62
(* experimental features *)
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    63
val ignore_no_atp =
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    64
  Attrib.setup_config_bool @{binding sledgehammer_ignore_no_atp} (K false)
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    65
val instantiate_inducts =
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    66
  Attrib.setup_config_bool @{binding sledgehammer_instantiate_inducts} (K false)
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    67
48292
7fcee834c7f5 more code rationalization in relevance filter
blanchet
parents: 48251
diff changeset
    68
val no_fact_override = {add = [], del = [], only = false}
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    69
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    70
fun needs_quoting reserved s =
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    71
  Symtab.defined reserved s orelse
50239
fb579401dc26 tuned signature;
wenzelm
parents: 50053
diff changeset
    72
  exists (not o Symbol_Pos.is_identifier) (Long_Name.explode s)
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    73
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    74
fun make_name reserved multi j name =
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    75
  (name |> needs_quoting reserved name ? quote) ^
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    76
  (if multi then "(" ^ string_of_int j ^ ")" else "")
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    77
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    78
fun explode_interval _ (Facts.FromTo (i, j)) = i upto j
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    79
  | explode_interval max (Facts.From i) = i upto i + max - 1
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    80
  | explode_interval _ (Facts.Single i) = [i]
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    81
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    82
val backquote =
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    83
  raw_explode #> map (fn "`" => "\\`" | s => s) #> implode #> enclose "`" "`"
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    84
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    85
(* unfolding these can yield really huge terms *)
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    86
val risky_defs = @{thms Bit0_def Bit1_def}
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    87
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    88
fun is_rec_eq lhs = Term.exists_subterm (curry (op =) (head_of lhs))
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    89
fun is_rec_def (@{const Trueprop} $ t) = is_rec_def t
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    90
  | is_rec_def (@{const ==>} $ _ $ t2) = is_rec_def t2
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    91
  | is_rec_def (Const (@{const_name "=="}, _) $ t1 $ t2) = is_rec_eq t1 t2
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    92
  | is_rec_def (Const (@{const_name HOL.eq}, _) $ t1 $ t2) = is_rec_eq t1 t2
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    93
  | is_rec_def _ = false
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    94
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    95
fun is_assum assms th = exists (fn ct => prop_of th aconv term_of ct) assms
48396
dd82d190c2af name tuning
blanchet
parents: 48394
diff changeset
    96
fun is_chained chained = member Thm.eq_thm_prop chained
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
    97
48396
dd82d190c2af name tuning
blanchet
parents: 48394
diff changeset
    98
fun scope_of_thm global assms chained th =
dd82d190c2af name tuning
blanchet
parents: 48394
diff changeset
    99
  if is_chained chained th then Chained
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   100
  else if global then Global
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   101
  else if is_assum assms th then Assum
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   102
  else Local
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   103
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   104
val may_be_induction =
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   105
  exists_subterm (fn Var (_, Type (@{type_name fun}, [_, T])) =>
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   106
                     body_type T = @{typ bool}
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   107
                   | _ => false)
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   108
53501
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   109
fun normalize_vars t =
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   110
  let
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   111
    fun normT (Type (s, Ts)) = fold_map normT Ts #>> curry Type s
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   112
      | normT (TVar (z as (_, S))) =
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   113
        (fn ((knownT, nT), accum) =>
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   114
            case find_index (equal z) knownT of
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   115
              ~1 => (TVar ((Name.uu, nT), S), ((z :: knownT, nT + 1), accum))
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   116
            | j => (TVar ((Name.uu, nT - j - 1), S), ((knownT, nT), accum)))
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   117
      | normT (T as TFree _) = pair T
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   118
    fun norm (t $ u) = norm t ##>> norm u #>> op $
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   119
      | norm (Const (s, T)) = normT T #>> curry Const s
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   120
      | norm (Var (z as (_, T))) =
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   121
        normT T
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   122
        #> (fn (T, (accumT, (known, n))) =>
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   123
               case find_index (equal z) known of
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   124
                 ~1 => (Var ((Name.uu, n), T), (accumT, (z :: known, n + 1)))
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   125
               | j => (Var ((Name.uu, n - j - 1), T), (accumT, (known, n))))
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   126
      | norm (Abs (_, T, t)) =
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   127
        norm t ##>> normT T #>> (fn (t, T) => Abs (Name.uu, T, t))
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   128
      | norm (Bound j) = pair (Bound j)
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   129
      | norm (Free (s, T)) = normT T #>> curry Free s
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   130
  in fst (norm t (([], 0), ([], 0))) end
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   131
48396
dd82d190c2af name tuning
blanchet
parents: 48394
diff changeset
   132
fun status_of_thm css name th =
53501
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   133
  let val t = normalize_vars (prop_of th) in
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   134
    (* FIXME: use structured name *)
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   135
    if String.isSubstring ".induct" name andalso may_be_induction t then
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   136
      Induction
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   137
    else case Termtab.lookup css t of
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   138
      SOME status => status
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   139
    | NONE =>
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   140
      let val concl = Logic.strip_imp_concl t in
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   141
        case try (HOLogic.dest_eq o HOLogic.dest_Trueprop) concl of
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   142
          SOME lrhss =>
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   143
          let
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   144
            val prems = Logic.strip_imp_prems t
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   145
            val t' = Logic.list_implies (prems, Logic.mk_equals lrhss)
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   146
          in
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   147
            Termtab.lookup css t' |> the_default General
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   148
          end
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   149
        | NONE => General
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   150
      end
b49bfa77958a speed up detection of simp rules
blanchet
parents: 52031
diff changeset
   151
  end
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   152
48396
dd82d190c2af name tuning
blanchet
parents: 48394
diff changeset
   153
fun stature_of_thm global assms chained css name th =
dd82d190c2af name tuning
blanchet
parents: 48394
diff changeset
   154
  (scope_of_thm global assms chained th, status_of_thm css name th)
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   155
52031
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51998
diff changeset
   156
fun fact_of_ref ctxt reserved chained css (xthm as (xref, args)) =
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   157
  let
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   158
    val ths = Attrib.eval_thms ctxt [xthm]
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   159
    val bracket =
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   160
      map (enclose "[" "]" o Pretty.str_of o Args.pretty_src ctxt) args
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   161
      |> implode
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   162
    fun nth_name j =
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   163
      case xref of
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   164
        Facts.Fact s => backquote s ^ bracket
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   165
      | Facts.Named (("", _), _) => "[" ^ bracket ^ "]"
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   166
      | Facts.Named ((name, _), NONE) =>
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   167
        make_name reserved (length ths > 1) (j + 1) name ^ bracket
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   168
      | Facts.Named ((name, _), SOME intervals) =>
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   169
        make_name reserved true
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   170
                 (nth (maps (explode_interval (length ths)) intervals) j) name ^
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   171
        bracket
48396
dd82d190c2af name tuning
blanchet
parents: 48394
diff changeset
   172
    fun add_nth th (j, rest) =
dd82d190c2af name tuning
blanchet
parents: 48394
diff changeset
   173
      let val name = nth_name j in
dd82d190c2af name tuning
blanchet
parents: 48394
diff changeset
   174
        (j + 1, ((name, stature_of_thm false [] chained css name th), th)
dd82d190c2af name tuning
blanchet
parents: 48394
diff changeset
   175
                :: rest)
dd82d190c2af name tuning
blanchet
parents: 48394
diff changeset
   176
      end
dd82d190c2af name tuning
blanchet
parents: 48394
diff changeset
   177
  in (0, []) |> fold add_nth ths |> snd end
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   178
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   179
(* Reject theorems with names like "List.filter.filter_list_def" or
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   180
  "Accessible_Part.acc.defs", as these are definitions arising from packages. *)
50736
07dfdf72ab75 tuned blacklisting in relevance filter
blanchet
parents: 50735
diff changeset
   181
fun is_package_def s =
07dfdf72ab75 tuned blacklisting in relevance filter
blanchet
parents: 50735
diff changeset
   182
  let val ss = Long_Name.explode s in
07dfdf72ab75 tuned blacklisting in relevance filter
blanchet
parents: 50735
diff changeset
   183
    length ss > 2 andalso not (hd ss = "local") andalso
07dfdf72ab75 tuned blacklisting in relevance filter
blanchet
parents: 50735
diff changeset
   184
    exists (fn suf => String.isSuffix suf s)
07dfdf72ab75 tuned blacklisting in relevance filter
blanchet
parents: 50735
diff changeset
   185
           ["_case_def", "_rec_def", "_size_def", "_size_overloaded_def"]
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   186
  end
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   187
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   188
(* FIXME: put other record thms here, or declare as "no_atp" *)
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   189
fun multi_base_blacklist ctxt ho_atp =
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   190
  ["defs", "select_defs", "update_defs", "split", "splits", "split_asm",
50442
4f6a4d32522c don't blacklist "case" theorems -- this causes problems in MaSh later
blanchet
parents: 50239
diff changeset
   191
   "ext_cases", "eq.simps", "eq.refl", "nchotomy", "case_cong",
51160
599ff65b85e2 systematic conversions between nat and nibble/char;
haftmann
parents: 51126
diff changeset
   192
   "weak_case_cong", "nat_of_char_simps", "nibble.simps",
48440
159e320ef851 identified "evil" theories for MaSh -- this is rather ad hoc, but so is MaSh anyway
blanchet
parents: 48438
diff changeset
   193
   "nibble.distinct"]
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   194
  |> not (ho_atp orelse (Config.get ctxt instantiate_inducts)) ?
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   195
        append ["induct", "inducts"]
48440
159e320ef851 identified "evil" theories for MaSh -- this is rather ad hoc, but so is MaSh anyway
blanchet
parents: 48438
diff changeset
   196
  |> map (prefix Long_Name.separator)
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   197
48667
ac58317ef11f rule out same "technical" theories for MePo as for MaSh
blanchet
parents: 48530
diff changeset
   198
(* FIXME: Ad hoc list *)
ac58317ef11f rule out same "technical" theories for MePo as for MaSh
blanchet
parents: 48530
diff changeset
   199
val technical_prefixes =
51126
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 51004
diff changeset
   200
  ["ATP", "Code_Evaluation", "Datatype", "Enum", "Lazy_Sequence",
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 51004
diff changeset
   201
   "Limited_Sequence", "Meson", "Metis", "Nitpick",
df86080de4cb reform of predicate compiler / quickcheck theories:
haftmann
parents: 51004
diff changeset
   202
   "Quickcheck_Random", "Quickcheck_Exhaustive", "Quickcheck_Narrowing",
50736
07dfdf72ab75 tuned blacklisting in relevance filter
blanchet
parents: 50735
diff changeset
   203
   "Random_Sequence", "Sledgehammer", "SMT"]
48667
ac58317ef11f rule out same "technical" theories for MePo as for MaSh
blanchet
parents: 48530
diff changeset
   204
  |> map (suffix Long_Name.separator)
48440
159e320ef851 identified "evil" theories for MaSh -- this is rather ad hoc, but so is MaSh anyway
blanchet
parents: 48438
diff changeset
   205
53507
a6ed27399ba9 avoid double traversal of term
blanchet
parents: 53506
diff changeset
   206
fun is_technical_const (s, _) =
48667
ac58317ef11f rule out same "technical" theories for MePo as for MaSh
blanchet
parents: 48530
diff changeset
   207
  exists (fn pref => String.isPrefix pref s) technical_prefixes
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   208
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   209
(* FIXME: make more reliable *)
53507
a6ed27399ba9 avoid double traversal of term
blanchet
parents: 53506
diff changeset
   210
val sep_class_sep = Long_Name.separator ^ "class" ^ Long_Name.separator
a6ed27399ba9 avoid double traversal of term
blanchet
parents: 53506
diff changeset
   211
fun is_low_level_class_const (s, _) =
a6ed27399ba9 avoid double traversal of term
blanchet
parents: 53506
diff changeset
   212
  s = @{const_name equal_class.equal} orelse String.isSubstring sep_class_sep s
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   213
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   214
fun is_that_fact th =
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   215
  String.isSuffix (Long_Name.separator ^ Obtain.thatN) (Thm.get_name_hint th)
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   216
  andalso exists_subterm (fn Free (s, _) => s = Name.skolem Auto_Bind.thesisN
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   217
                           | _ => false) (prop_of th)
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   218
50523
0799339fea0f get rid of some junk facts in the MaSh evaluation driver
blanchet
parents: 50512
diff changeset
   219
fun is_likely_tautology_too_meta_or_too_technical th =
48406
b002cc16aa99 honor suggested MaSh weights
blanchet
parents: 48400
diff changeset
   220
  let
50053
fea589c8583e fixed detection of tautologies -- things like "a = b" in a structured proof, where a and b are Frees, shouldn't be discarted as tautologies
blanchet
parents: 50049
diff changeset
   221
    fun is_interesting_subterm (Const (s, _)) =
fea589c8583e fixed detection of tautologies -- things like "a = b" in a structured proof, where a and b are Frees, shouldn't be discarted as tautologies
blanchet
parents: 50049
diff changeset
   222
        not (member (op =) atp_widely_irrelevant_consts s)
fea589c8583e fixed detection of tautologies -- things like "a = b" in a structured proof, where a and b are Frees, shouldn't be discarted as tautologies
blanchet
parents: 50049
diff changeset
   223
      | is_interesting_subterm (Free _) = true
fea589c8583e fixed detection of tautologies -- things like "a = b" in a structured proof, where a and b are Frees, shouldn't be discarted as tautologies
blanchet
parents: 50049
diff changeset
   224
      | is_interesting_subterm _ = false
48406
b002cc16aa99 honor suggested MaSh weights
blanchet
parents: 48400
diff changeset
   225
    fun is_boring_bool t =
50053
fea589c8583e fixed detection of tautologies -- things like "a = b" in a structured proof, where a and b are Frees, shouldn't be discarted as tautologies
blanchet
parents: 50049
diff changeset
   226
      not (exists_subterm is_interesting_subterm t) orelse
48406
b002cc16aa99 honor suggested MaSh weights
blanchet
parents: 48400
diff changeset
   227
      exists_type (exists_subtype (curry (op =) @{typ prop})) t
50495
bd9a0028b063 better tautology check -- don't reject "prod_cases3" for example
blanchet
parents: 50490
diff changeset
   228
    fun is_boring_prop _ (@{const Trueprop} $ t) = is_boring_bool t
bd9a0028b063 better tautology check -- don't reject "prod_cases3" for example
blanchet
parents: 50490
diff changeset
   229
      | is_boring_prop Ts (@{const "==>"} $ t $ u) =
bd9a0028b063 better tautology check -- don't reject "prod_cases3" for example
blanchet
parents: 50490
diff changeset
   230
        is_boring_prop Ts t andalso is_boring_prop Ts u
50496
8665ec681e47 further fix related to bd9a0028b063 -- that change was per se right, but it exposed a bug in the pattern for "all"
blanchet
parents: 50495
diff changeset
   231
      | is_boring_prop Ts (Const (@{const_name all}, _) $ Abs (_, T, t)) =
8665ec681e47 further fix related to bd9a0028b063 -- that change was per se right, but it exposed a bug in the pattern for "all"
blanchet
parents: 50495
diff changeset
   232
        is_boring_prop (T :: Ts) t
50495
bd9a0028b063 better tautology check -- don't reject "prod_cases3" for example
blanchet
parents: 50490
diff changeset
   233
      | is_boring_prop Ts ((t as Const (@{const_name all}, _)) $ u) =
bd9a0028b063 better tautology check -- don't reject "prod_cases3" for example
blanchet
parents: 50490
diff changeset
   234
        is_boring_prop Ts (t $ eta_expand Ts u 1)
bd9a0028b063 better tautology check -- don't reject "prod_cases3" for example
blanchet
parents: 50490
diff changeset
   235
      | is_boring_prop _ (Const (@{const_name "=="}, _) $ t $ u) =
48406
b002cc16aa99 honor suggested MaSh weights
blanchet
parents: 48400
diff changeset
   236
        is_boring_bool t andalso is_boring_bool u
50495
bd9a0028b063 better tautology check -- don't reject "prod_cases3" for example
blanchet
parents: 50490
diff changeset
   237
      | is_boring_prop _ _ = true
50523
0799339fea0f get rid of some junk facts in the MaSh evaluation driver
blanchet
parents: 50512
diff changeset
   238
    val t = prop_of th
48406
b002cc16aa99 honor suggested MaSh weights
blanchet
parents: 48400
diff changeset
   239
  in
53508
8d8f72aa5c0b got rid of another taboo that appears to make no difference in practice (and that slows down the relevance filter)
blanchet
parents: 53507
diff changeset
   240
    (is_boring_prop [] t andalso not (Thm.eq_thm_prop (@{thm ext}, th))) orelse
53507
a6ed27399ba9 avoid double traversal of term
blanchet
parents: 53506
diff changeset
   241
    exists_Const (is_technical_const orf is_low_level_class_const) t orelse
a6ed27399ba9 avoid double traversal of term
blanchet
parents: 53506
diff changeset
   242
    is_that_fact th
48438
3e45c98fe127 distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
blanchet
parents: 48437
diff changeset
   243
  end
3e45c98fe127 distinguish between recursive and nonrecursive definitions + clean up typedef dependencies in MaSh
blanchet
parents: 48437
diff changeset
   244
50511
8825c36cb1ce don't query blacklisted theorems in evaluation driver
blanchet
parents: 50510
diff changeset
   245
fun is_blacklisted_or_something ctxt ho_atp name =
50510
7e4f2f8d9b50 export a pair of ML functions
blanchet
parents: 50496
diff changeset
   246
  (not (Config.get ctxt ignore_no_atp) andalso is_package_def name) orelse
7e4f2f8d9b50 export a pair of ML functions
blanchet
parents: 50496
diff changeset
   247
  exists (fn s => String.isSuffix s name) (multi_base_blacklist ctxt ho_atp)
7e4f2f8d9b50 export a pair of ML functions
blanchet
parents: 50496
diff changeset
   248
51998
f732a674db1b renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents: 51160
diff changeset
   249
fun hackish_string_of_term ctxt =
50049
dd6a4655cd72 avoid messing too much with output of "string_of_term", so that it doesn't break the yxml encoding for jEdit
blanchet
parents: 50048
diff changeset
   250
  with_vanilla_print_mode (Syntax.string_of_term ctxt) #> simplify_spaces
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   251
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   252
(* This is a terrible hack. Free variables are sometimes coded as "M__" when
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   253
   they are displayed as "M" and we want to avoid clashes with these. But
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   254
   sometimes it's even worse: "Ma__" encodes "M". So we simply reserve all
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   255
   prefixes of all free variables. In the worse case scenario, where the fact
49981
blanchet
parents: 48667
diff changeset
   256
   won't be resolved correctly, the user can fix it manually, e.g., by giving a
blanchet
parents: 48667
diff changeset
   257
   name to the offending fact. *)
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   258
fun all_prefixes_of s =
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   259
  map (fn i => String.extract (s, 0, SOME i)) (1 upto size s - 1)
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   260
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   261
fun close_form t =
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   262
  (t, [] |> Term.add_free_names t |> maps all_prefixes_of)
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   263
  |> fold (fn ((s, i), T) => fn (t', taken) =>
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   264
              let val s' = singleton (Name.variant_list taken) s in
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   265
                ((if fastype_of t' = HOLogic.boolT then HOLogic.all_const
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   266
                  else Logic.all_const) T
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   267
                 $ Abs (s', T, abstract_over (Var ((s, i), T), t')),
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   268
                 s' :: taken)
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   269
              end)
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   270
          (Term.add_vars t [] |> sort_wrt (fst o fst))
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   271
  |> fst
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   272
51998
f732a674db1b renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents: 51160
diff changeset
   273
fun backquote_term ctxt = close_form #> hackish_string_of_term ctxt #> backquote
f732a674db1b renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents: 51160
diff changeset
   274
fun backquote_thm ctxt = backquote_term ctxt o prop_of
48394
82fc8c956cdc fixed various issues with MaSh's file handling + tune output + generate local facts again + handle nameless facts gracefully
blanchet
parents: 48332
diff changeset
   275
53502
2eaaca796234 gracefully handle huge thys
blanchet
parents: 53501
diff changeset
   276
(* gracefully handle huge background theories *)
2eaaca796234 gracefully handle huge thys
blanchet
parents: 53501
diff changeset
   277
val max_simps_for_clasimpset = 10000
2eaaca796234 gracefully handle huge thys
blanchet
parents: 53501
diff changeset
   278
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   279
fun clasimpset_rule_table_of ctxt =
53502
2eaaca796234 gracefully handle huge thys
blanchet
parents: 53501
diff changeset
   280
  let val simps = ctxt |> simpset_of |> dest_ss |> #simps in
2eaaca796234 gracefully handle huge thys
blanchet
parents: 53501
diff changeset
   281
    if length simps >= max_simps_for_clasimpset then
2eaaca796234 gracefully handle huge thys
blanchet
parents: 53501
diff changeset
   282
      Termtab.empty
2eaaca796234 gracefully handle huge thys
blanchet
parents: 53501
diff changeset
   283
    else
2eaaca796234 gracefully handle huge thys
blanchet
parents: 53501
diff changeset
   284
      let
2eaaca796234 gracefully handle huge thys
blanchet
parents: 53501
diff changeset
   285
        fun add stature th =
2eaaca796234 gracefully handle huge thys
blanchet
parents: 53501
diff changeset
   286
          Termtab.update (normalize_vars (prop_of th), stature)
2eaaca796234 gracefully handle huge thys
blanchet
parents: 53501
diff changeset
   287
        val {safeIs, (* safeEs, *) hazIs, (* hazEs, *) ...} =
2eaaca796234 gracefully handle huge thys
blanchet
parents: 53501
diff changeset
   288
          ctxt |> claset_of |> Classical.rep_cs
2eaaca796234 gracefully handle huge thys
blanchet
parents: 53501
diff changeset
   289
        val intros = Item_Net.content safeIs @ Item_Net.content hazIs
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   290
(* Add once it is used:
53502
2eaaca796234 gracefully handle huge thys
blanchet
parents: 53501
diff changeset
   291
        val elims =
2eaaca796234 gracefully handle huge thys
blanchet
parents: 53501
diff changeset
   292
          Item_Net.content safeEs @ Item_Net.content hazEs
2eaaca796234 gracefully handle huge thys
blanchet
parents: 53501
diff changeset
   293
          |> map Classical.classical_rule
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   294
*)
53502
2eaaca796234 gracefully handle huge thys
blanchet
parents: 53501
diff changeset
   295
        val specs = ctxt |> Spec_Rules.get
2eaaca796234 gracefully handle huge thys
blanchet
parents: 53501
diff changeset
   296
        val (rec_defs, nonrec_defs) =
2eaaca796234 gracefully handle huge thys
blanchet
parents: 53501
diff changeset
   297
          specs |> filter (curry (op =) Spec_Rules.Equational o fst)
2eaaca796234 gracefully handle huge thys
blanchet
parents: 53501
diff changeset
   298
                |> maps (snd o snd)
2eaaca796234 gracefully handle huge thys
blanchet
parents: 53501
diff changeset
   299
                |> filter_out (member Thm.eq_thm_prop risky_defs)
2eaaca796234 gracefully handle huge thys
blanchet
parents: 53501
diff changeset
   300
                |> List.partition (is_rec_def o prop_of)
2eaaca796234 gracefully handle huge thys
blanchet
parents: 53501
diff changeset
   301
        val spec_intros =
2eaaca796234 gracefully handle huge thys
blanchet
parents: 53501
diff changeset
   302
          specs |> filter (member (op =) [Spec_Rules.Inductive,
2eaaca796234 gracefully handle huge thys
blanchet
parents: 53501
diff changeset
   303
                                          Spec_Rules.Co_Inductive] o fst)
2eaaca796234 gracefully handle huge thys
blanchet
parents: 53501
diff changeset
   304
                |> maps (snd o snd)
2eaaca796234 gracefully handle huge thys
blanchet
parents: 53501
diff changeset
   305
      in
2eaaca796234 gracefully handle huge thys
blanchet
parents: 53501
diff changeset
   306
        Termtab.empty |> fold (add Simp o snd) simps
2eaaca796234 gracefully handle huge thys
blanchet
parents: 53501
diff changeset
   307
                      |> fold (add Rec_Def) rec_defs
2eaaca796234 gracefully handle huge thys
blanchet
parents: 53501
diff changeset
   308
                      |> fold (add Non_Rec_Def) nonrec_defs
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   309
(* Add once it is used:
53502
2eaaca796234 gracefully handle huge thys
blanchet
parents: 53501
diff changeset
   310
                      |> fold (add Elim) elims
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   311
*)
53502
2eaaca796234 gracefully handle huge thys
blanchet
parents: 53501
diff changeset
   312
                      |> fold (add Intro) intros
2eaaca796234 gracefully handle huge thys
blanchet
parents: 53501
diff changeset
   313
                      |> fold (add Inductive) spec_intros
2eaaca796234 gracefully handle huge thys
blanchet
parents: 53501
diff changeset
   314
      end
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   315
  end
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   316
50481
5d147d492792 push normalization further -- avoid theorems that are duplicates of each other except for equality symmetry (esp. for "list.distinct(1)" vs. "(2)" etc.)
blanchet
parents: 50442
diff changeset
   317
fun normalize_eq (t as @{const Trueprop}
5d147d492792 push normalization further -- avoid theorems that are duplicates of each other except for equality symmetry (esp. for "list.distinct(1)" vs. "(2)" etc.)
blanchet
parents: 50442
diff changeset
   318
        $ ((t0 as Const (@{const_name HOL.eq}, _)) $ t1 $ t2)) =
5d147d492792 push normalization further -- avoid theorems that are duplicates of each other except for equality symmetry (esp. for "list.distinct(1)" vs. "(2)" etc.)
blanchet
parents: 50442
diff changeset
   319
    if Term_Ord.fast_term_ord (t1, t2) <> GREATER then t
5d147d492792 push normalization further -- avoid theorems that are duplicates of each other except for equality symmetry (esp. for "list.distinct(1)" vs. "(2)" etc.)
blanchet
parents: 50442
diff changeset
   320
    else HOLogic.mk_Trueprop (t0 $ t2 $ t1)
5d147d492792 push normalization further -- avoid theorems that are duplicates of each other except for equality symmetry (esp. for "list.distinct(1)" vs. "(2)" etc.)
blanchet
parents: 50442
diff changeset
   321
  | normalize_eq (t as @{const Trueprop} $ (@{const Not}
5d147d492792 push normalization further -- avoid theorems that are duplicates of each other except for equality symmetry (esp. for "list.distinct(1)" vs. "(2)" etc.)
blanchet
parents: 50442
diff changeset
   322
        $ ((t0 as Const (@{const_name HOL.eq}, _)) $ t1 $ t2))) =
5d147d492792 push normalization further -- avoid theorems that are duplicates of each other except for equality symmetry (esp. for "list.distinct(1)" vs. "(2)" etc.)
blanchet
parents: 50442
diff changeset
   323
    if Term_Ord.fast_term_ord (t1, t2) <> GREATER then t
5d147d492792 push normalization further -- avoid theorems that are duplicates of each other except for equality symmetry (esp. for "list.distinct(1)" vs. "(2)" etc.)
blanchet
parents: 50442
diff changeset
   324
    else HOLogic.mk_Trueprop (HOLogic.mk_not (t0 $ t2 $ t1))
5d147d492792 push normalization further -- avoid theorems that are duplicates of each other except for equality symmetry (esp. for "list.distinct(1)" vs. "(2)" etc.)
blanchet
parents: 50442
diff changeset
   325
  | normalize_eq t = t
5d147d492792 push normalization further -- avoid theorems that are duplicates of each other except for equality symmetry (esp. for "list.distinct(1)" vs. "(2)" etc.)
blanchet
parents: 50442
diff changeset
   326
53503
d2f21e305d0c stronger fact normalization
blanchet
parents: 53502
diff changeset
   327
val normalize_eq_vars = normalize_eq #> normalize_vars
50485
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50481
diff changeset
   328
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50481
diff changeset
   329
fun if_thm_before th th' =
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50481
diff changeset
   330
  if Theory.subthy (pairself Thm.theory_of_thm (th, th')) then th else th'
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50481
diff changeset
   331
50734
73612ad116e6 learn from low-level, inside-class facts
blanchet
parents: 50733
diff changeset
   332
(* Hack: Conflate the facts about a class as seen from the outside with the
73612ad116e6 learn from low-level, inside-class facts
blanchet
parents: 50733
diff changeset
   333
   corresponding low-level facts, so that MaSh can learn from the low-level
73612ad116e6 learn from low-level, inside-class facts
blanchet
parents: 50733
diff changeset
   334
   proofs. *)
73612ad116e6 learn from low-level, inside-class facts
blanchet
parents: 50733
diff changeset
   335
fun un_class_ify s =
73612ad116e6 learn from low-level, inside-class facts
blanchet
parents: 50733
diff changeset
   336
  case first_field "_class" s of
73612ad116e6 learn from low-level, inside-class facts
blanchet
parents: 50733
diff changeset
   337
    SOME (pref, suf) => [s, pref ^ suf]
73612ad116e6 learn from low-level, inside-class facts
blanchet
parents: 50733
diff changeset
   338
  | NONE => [s]
73612ad116e6 learn from low-level, inside-class facts
blanchet
parents: 50733
diff changeset
   339
50735
6b232d76cbc9 refined class handling, to prevent cycles in fact graph
blanchet
parents: 50734
diff changeset
   340
fun build_name_tables name_of facts =
50733
blanchet
parents: 50586
diff changeset
   341
  let
53503
d2f21e305d0c stronger fact normalization
blanchet
parents: 53502
diff changeset
   342
    fun cons_thm (_, th) =
d2f21e305d0c stronger fact normalization
blanchet
parents: 53502
diff changeset
   343
      Termtab.cons_list (normalize_eq_vars (prop_of th), th)
50735
6b232d76cbc9 refined class handling, to prevent cycles in fact graph
blanchet
parents: 50734
diff changeset
   344
    fun add_plain canon alias =
50815
41b804049fae make name table work the way it was intended to
blanchet
parents: 50756
diff changeset
   345
      Symtab.update (Thm.get_name_hint alias,
50735
6b232d76cbc9 refined class handling, to prevent cycles in fact graph
blanchet
parents: 50734
diff changeset
   346
                     name_of (if_thm_before canon alias))
6b232d76cbc9 refined class handling, to prevent cycles in fact graph
blanchet
parents: 50734
diff changeset
   347
    fun add_plains (_, aliases as canon :: _) = fold (add_plain canon) aliases
6b232d76cbc9 refined class handling, to prevent cycles in fact graph
blanchet
parents: 50734
diff changeset
   348
    fun add_inclass (name, target) =
6b232d76cbc9 refined class handling, to prevent cycles in fact graph
blanchet
parents: 50734
diff changeset
   349
      fold (fn s => Symtab.update (s, target)) (un_class_ify name)
50815
41b804049fae make name table work the way it was intended to
blanchet
parents: 50756
diff changeset
   350
    val prop_tab = fold cons_thm facts Termtab.empty
50735
6b232d76cbc9 refined class handling, to prevent cycles in fact graph
blanchet
parents: 50734
diff changeset
   351
    val plain_name_tab = Termtab.fold add_plains prop_tab Symtab.empty
6b232d76cbc9 refined class handling, to prevent cycles in fact graph
blanchet
parents: 50734
diff changeset
   352
    val inclass_name_tab = Symtab.fold add_inclass plain_name_tab Symtab.empty
6b232d76cbc9 refined class handling, to prevent cycles in fact graph
blanchet
parents: 50734
diff changeset
   353
  in (plain_name_tab, inclass_name_tab) end
50485
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50481
diff changeset
   354
53504
9750618c32ed faster uniquification
blanchet
parents: 53503
diff changeset
   355
fun keyed_distinct key_of xs =
9750618c32ed faster uniquification
blanchet
parents: 53503
diff changeset
   356
  let val tab = fold (Termtab.default o `key_of) xs Termtab.empty in
9750618c32ed faster uniquification
blanchet
parents: 53503
diff changeset
   357
    Termtab.fold (cons o snd) tab []
9750618c32ed faster uniquification
blanchet
parents: 53503
diff changeset
   358
  end
9750618c32ed faster uniquification
blanchet
parents: 53503
diff changeset
   359
9750618c32ed faster uniquification
blanchet
parents: 53503
diff changeset
   360
fun hashed_keyed_distinct hash_of key_of xs =
9750618c32ed faster uniquification
blanchet
parents: 53503
diff changeset
   361
  let
9750618c32ed faster uniquification
blanchet
parents: 53503
diff changeset
   362
    val ys = map (`hash_of) xs
9750618c32ed faster uniquification
blanchet
parents: 53503
diff changeset
   363
    val sorted_ys = sort (int_ord o pairself fst) ys
9750618c32ed faster uniquification
blanchet
parents: 53503
diff changeset
   364
    val grouped_ys = AList.coalesce (op =) sorted_ys
9750618c32ed faster uniquification
blanchet
parents: 53503
diff changeset
   365
  in maps (keyed_distinct key_of o snd) grouped_ys end
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   366
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   367
fun struct_induct_rule_on th =
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   368
  case Logic.strip_horn (prop_of th) of
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   369
    (prems, @{const Trueprop}
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   370
            $ ((p as Var ((p_name, 0), _)) $ (a as Var (_, ind_T)))) =>
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   371
    if not (is_TVar ind_T) andalso length prems > 1 andalso
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   372
       exists (exists_subterm (curry (op aconv) p)) prems andalso
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   373
       not (exists (exists_subterm (curry (op aconv) a)) prems) then
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   374
      SOME (p_name, ind_T)
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   375
    else
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   376
      NONE
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   377
  | _ => NONE
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   378
50586
e31ee2076db1 add a timeout in induction rule instantiation
blanchet
parents: 50523
diff changeset
   379
val instantiate_induct_timeout = seconds 0.01
e31ee2076db1 add a timeout in induction rule instantiation
blanchet
parents: 50523
diff changeset
   380
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   381
fun instantiate_induct_rule ctxt concl_prop p_name ((name, stature), th) ind_x =
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   382
  let
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   383
    fun varify_noninducts (t as Free (s, T)) =
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   384
        if (s, T) = ind_x orelse can dest_funT T then t else Var ((s, 0), T)
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   385
      | varify_noninducts t = t
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   386
    val p_inst =
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   387
      concl_prop |> map_aterms varify_noninducts |> close_form
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   388
                 |> lambda (Free ind_x)
51998
f732a674db1b renamed Sledgehammer functions with 'for' in their names to 'of'
blanchet
parents: 51160
diff changeset
   389
                 |> hackish_string_of_term ctxt
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   390
  in
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   391
    ((fn () => name () ^ "[where " ^ p_name ^ " = " ^ quote p_inst ^ "]",
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   392
      stature), th |> read_instantiate ctxt [((p_name, 0), p_inst)])
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   393
  end
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   394
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   395
fun type_match thy (T1, T2) =
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   396
  (Sign.typ_match thy (T2, T1) Vartab.empty; true)
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   397
  handle Type.TYPE_MATCH => false
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   398
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   399
fun instantiate_if_induct_rule ctxt stmt stmt_xs (ax as (_, th)) =
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   400
  case struct_induct_rule_on th of
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   401
    SOME (p_name, ind_T) =>
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   402
    let val thy = Proof_Context.theory_of ctxt in
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   403
      stmt_xs |> filter (fn (_, T) => type_match thy (T, ind_T))
50586
e31ee2076db1 add a timeout in induction rule instantiation
blanchet
parents: 50523
diff changeset
   404
              |> map_filter (try (TimeLimit.timeLimit instantiate_induct_timeout
e31ee2076db1 add a timeout in induction rule instantiation
blanchet
parents: 50523
diff changeset
   405
                     (instantiate_induct_rule ctxt stmt p_name ax)))
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   406
    end
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   407
  | NONE => [ax]
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   408
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   409
fun external_frees t =
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   410
  [] |> Term.add_frees t |> filter_out (can Name.dest_internal o fst)
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   411
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   412
fun maybe_instantiate_inducts ctxt hyp_ts concl_t =
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   413
  if Config.get ctxt instantiate_inducts then
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   414
    let
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   415
      val thy = Proof_Context.theory_of ctxt
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   416
      val ind_stmt =
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   417
        (hyp_ts |> filter_out (null o external_frees), concl_t)
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   418
        |> Logic.list_implies |> Object_Logic.atomize_term thy
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   419
      val ind_stmt_xs = external_frees ind_stmt
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   420
    in maps (instantiate_if_induct_rule ctxt ind_stmt ind_stmt_xs) end
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   421
  else
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   422
    I
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   423
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   424
fun maybe_filter_no_atps ctxt =
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   425
  not (Config.get ctxt ignore_no_atp) ? filter_out (No_ATPs.member ctxt o snd)
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   426
51004
5f2788c38127 distinguish raw and non-raw facts, using raw for 10 000s of facts and non-raw after selection of some hundreds
blanchet
parents: 50815
diff changeset
   427
fun fact_of_raw_fact ((name, stature), th) = ((name (), stature), th)
5f2788c38127 distinguish raw and non-raw facts, using raw for 10 000s of facts and non-raw after selection of some hundreds
blanchet
parents: 50815
diff changeset
   428
50523
0799339fea0f get rid of some junk facts in the MaSh evaluation driver
blanchet
parents: 50512
diff changeset
   429
fun all_facts ctxt generous ho_atp reserved add_ths chained css =
48251
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 48250
diff changeset
   430
  let
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 48250
diff changeset
   431
    val thy = Proof_Context.theory_of ctxt
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 48250
diff changeset
   432
    val global_facts = Global_Theory.facts_of thy
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 48250
diff changeset
   433
    val local_facts = Proof_Context.facts_of ctxt
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 48250
diff changeset
   434
    val named_locals = local_facts |> Facts.dest_static []
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 48250
diff changeset
   435
    val assms = Assumption.all_assms_of ctxt
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 48250
diff changeset
   436
    fun is_good_unnamed_local th =
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 48250
diff changeset
   437
      not (Thm.has_name_hint th) andalso
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 48250
diff changeset
   438
      forall (fn (_, ths) => not (member Thm.eq_thm_prop ths th)) named_locals
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 48250
diff changeset
   439
    val unnamed_locals =
48396
dd82d190c2af name tuning
blanchet
parents: 48394
diff changeset
   440
      union Thm.eq_thm_prop (Facts.props local_facts) chained
48251
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 48250
diff changeset
   441
      |> filter is_good_unnamed_local |> map (pair "" o single)
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 48250
diff changeset
   442
    val full_space =
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 48250
diff changeset
   443
      Name_Space.merge (Facts.space_of global_facts, Facts.space_of local_facts)
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 48250
diff changeset
   444
    fun add_facts global foldx facts =
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 48250
diff changeset
   445
      foldx (fn (name0, ths) =>
50512
c283bc0a8f1a tweaked which facts are included for MaSh evaluations
blanchet
parents: 50511
diff changeset
   446
        if name0 <> "" andalso
48251
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 48250
diff changeset
   447
           forall (not o member Thm.eq_thm_prop add_ths) ths andalso
50511
8825c36cb1ce don't query blacklisted theorems in evaluation driver
blanchet
parents: 50510
diff changeset
   448
           (Facts.is_concealed facts name0 orelse
8825c36cb1ce don't query blacklisted theorems in evaluation driver
blanchet
parents: 50510
diff changeset
   449
            not (can (Proof_Context.get_thms ctxt) name0) orelse
50523
0799339fea0f get rid of some junk facts in the MaSh evaluation driver
blanchet
parents: 50512
diff changeset
   450
            (not generous andalso
50512
c283bc0a8f1a tweaked which facts are included for MaSh evaluations
blanchet
parents: 50511
diff changeset
   451
             is_blacklisted_or_something ctxt ho_atp name0)) then
48251
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 48250
diff changeset
   452
          I
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 48250
diff changeset
   453
        else
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 48250
diff changeset
   454
          let
50485
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50481
diff changeset
   455
            val n = length ths
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50481
diff changeset
   456
            val multi = n > 1
48251
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 48250
diff changeset
   457
            fun check_thms a =
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 48250
diff changeset
   458
              case try (Proof_Context.get_thms ctxt) a of
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 48250
diff changeset
   459
                NONE => false
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 48250
diff changeset
   460
              | SOME ths' => eq_list Thm.eq_thm_prop (ths, ths')
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 48250
diff changeset
   461
          in
50485
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50481
diff changeset
   462
            pair n
50756
c96bb430ddb0 put single-theorem names before multi-theorem ones (broken since 5d147d492792)
blanchet
parents: 50736
diff changeset
   463
            #> fold_rev (fn th => fn (j, accum) =>
50485
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50481
diff changeset
   464
                   (j - 1,
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50481
diff changeset
   465
                    if not (member Thm.eq_thm_prop add_ths th) andalso
53506
f99ee3adb81d got rid of old, needless logic
blanchet
parents: 53504
diff changeset
   466
                       (is_likely_tautology_too_meta_or_too_technical th) then
50756
c96bb430ddb0 put single-theorem names before multi-theorem ones (broken since 5d147d492792)
blanchet
parents: 50736
diff changeset
   467
                      accum
50485
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50481
diff changeset
   468
                    else
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50481
diff changeset
   469
                      let
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50481
diff changeset
   470
                        val new =
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50481
diff changeset
   471
                          (((fn () =>
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50481
diff changeset
   472
                                if name0 = "" then
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50481
diff changeset
   473
                                  backquote_thm ctxt th
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50481
diff changeset
   474
                                else
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50481
diff changeset
   475
                                  [Facts.extern ctxt facts name0,
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50481
diff changeset
   476
                                   Name_Space.extern ctxt full_space name0]
50756
c96bb430ddb0 put single-theorem names before multi-theorem ones (broken since 5d147d492792)
blanchet
parents: 50736
diff changeset
   477
                                  |> distinct (op =)
50485
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50481
diff changeset
   478
                                  |> find_first check_thms
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50481
diff changeset
   479
                                  |> the_default name0
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50481
diff changeset
   480
                                  |> make_name reserved multi j),
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50481
diff changeset
   481
                             stature_of_thm global assms chained css name0 th),
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50481
diff changeset
   482
                           th)
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50481
diff changeset
   483
                      in
50756
c96bb430ddb0 put single-theorem names before multi-theorem ones (broken since 5d147d492792)
blanchet
parents: 50736
diff changeset
   484
                        accum |> (if multi then apsnd else apfst) (cons new)
50485
3c6ac2da2f45 merge aliased theorems in MaSh dependencies, modulo symmetry of equality
blanchet
parents: 50481
diff changeset
   485
                      end)) ths
48251
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 48250
diff changeset
   486
            #> snd
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 48250
diff changeset
   487
          end)
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 48250
diff changeset
   488
  in
50756
c96bb430ddb0 put single-theorem names before multi-theorem ones (broken since 5d147d492792)
blanchet
parents: 50736
diff changeset
   489
    (* The single-theorem names go before the multiple-theorem ones (e.g.,
c96bb430ddb0 put single-theorem names before multi-theorem ones (broken since 5d147d492792)
blanchet
parents: 50736
diff changeset
   490
       "xxx" vs. "xxx(3)"), so that single names are preferred when both are
c96bb430ddb0 put single-theorem names before multi-theorem ones (broken since 5d147d492792)
blanchet
parents: 50736
diff changeset
   491
       available. *)
c96bb430ddb0 put single-theorem names before multi-theorem ones (broken since 5d147d492792)
blanchet
parents: 50736
diff changeset
   492
    `I [] |> add_facts false fold local_facts (unnamed_locals @ named_locals)
c96bb430ddb0 put single-theorem names before multi-theorem ones (broken since 5d147d492792)
blanchet
parents: 50736
diff changeset
   493
          |> add_facts true Facts.fold_static global_facts global_facts
c96bb430ddb0 put single-theorem names before multi-theorem ones (broken since 5d147d492792)
blanchet
parents: 50736
diff changeset
   494
          |> op @
48251
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 48250
diff changeset
   495
  end
6cdcfbddc077 moved most of MaSh exporter code to Sledgehammer
blanchet
parents: 48250
diff changeset
   496
48396
dd82d190c2af name tuning
blanchet
parents: 48394
diff changeset
   497
fun nearly_all_facts ctxt ho_atp {add, del, only} reserved css chained hyp_ts
dd82d190c2af name tuning
blanchet
parents: 48394
diff changeset
   498
                     concl_t =
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   499
  if only andalso null add then
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   500
    []
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   501
  else
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   502
    let
48396
dd82d190c2af name tuning
blanchet
parents: 48394
diff changeset
   503
      val chained =
dd82d190c2af name tuning
blanchet
parents: 48394
diff changeset
   504
        chained
48292
7fcee834c7f5 more code rationalization in relevance filter
blanchet
parents: 48251
diff changeset
   505
        |> maps (fn th => insert Thm.eq_thm_prop (zero_var_indexes th) [th])
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   506
    in
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   507
      (if only then
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   508
         maps (map (fn ((name, stature), th) => ((K name, stature), th))
52031
9a9238342963 tuning -- renamed '_from_' to '_of_' in Sledgehammer
blanchet
parents: 51998
diff changeset
   509
               o fact_of_ref ctxt reserved chained css) add
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   510
       else
48292
7fcee834c7f5 more code rationalization in relevance filter
blanchet
parents: 48251
diff changeset
   511
         let val (add, del) = pairself (Attrib.eval_thms ctxt) (add, del) in
50442
4f6a4d32522c don't blacklist "case" theorems -- this causes problems in MaSh later
blanchet
parents: 50239
diff changeset
   512
           all_facts ctxt false ho_atp reserved add chained css
48292
7fcee834c7f5 more code rationalization in relevance filter
blanchet
parents: 48251
diff changeset
   513
           |> filter_out (member Thm.eq_thm_prop del o snd)
7fcee834c7f5 more code rationalization in relevance filter
blanchet
parents: 48251
diff changeset
   514
           |> maybe_filter_no_atps ctxt
53504
9750618c32ed faster uniquification
blanchet
parents: 53503
diff changeset
   515
           |> hashed_keyed_distinct (size_of_term o prop_of o snd)
9750618c32ed faster uniquification
blanchet
parents: 53503
diff changeset
   516
                  (normalize_eq_vars o prop_of o snd)
48292
7fcee834c7f5 more code rationalization in relevance filter
blanchet
parents: 48251
diff changeset
   517
         end)
48250
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   518
      |> maybe_instantiate_inducts ctxt hyp_ts concl_t
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   519
    end
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   520
1065c307fafe further ML structure split to permit finer-grained loading/reordering (problem to solve: MaSh needs most of Sledgehammer)
blanchet
parents:
diff changeset
   521
end;