src/Pure/Tools/find_theorems.ML
author wenzelm
Fri, 16 Aug 2019 10:33:25 +0200
changeset 70545 b93ba98e627a
parent 67721 5348bea4accd
child 74525 c960bfcb91db
permissions -rw-r--r--
tuned signature;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
30143
98a986b02022 observe basic Isabelle/ML coding conventions;
wenzelm
parents: 30142
diff changeset
     1
(*  Title:      Pure/Tools/find_theorems.ML
26283
e19a5a7e83f1 more precise Author line;
wenzelm
parents: 25992
diff changeset
     2
    Author:     Rafal Kolanski and Gerwin Klein, NICTA
46718
wenzelm
parents: 46717
diff changeset
     3
    Author:     Lars Noschinski and Alexander Krauss, TU Muenchen
16033
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
     4
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
     5
Retrieve theorems from proof context.
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
     6
*)
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
     7
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
     8
signature FIND_THEOREMS =
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
     9
sig
16036
1da07ac33711 added read_criterion/pretty_criterion;
wenzelm
parents: 16033
diff changeset
    10
  datatype 'term criterion =
46717
b09afce1e54f removed broken/unused introiff (cf. d452117ba564);
wenzelm
parents: 46716
diff changeset
    11
    Name of string | Intro | Elim | Dest | Solves | Simp of 'term | Pattern of 'term
43070
0318781be055 explicit type for search queries
krauss
parents: 43069
diff changeset
    12
  type 'term query = {
0318781be055 explicit type for search queries
krauss
parents: 43069
diff changeset
    13
    goal: thm option,
0318781be055 explicit type for search queries
krauss
parents: 43069
diff changeset
    14
    limit: int option,
0318781be055 explicit type for search queries
krauss
parents: 43069
diff changeset
    15
    rem_dups: bool,
0318781be055 explicit type for search queries
krauss
parents: 43069
diff changeset
    16
    criteria: (bool * 'term criterion) list
0318781be055 explicit type for search queries
krauss
parents: 43069
diff changeset
    17
  }
62848
e4140efe699e clarified bootstrap -- more uniform use of ML files;
wenzelm
parents: 61841
diff changeset
    18
  val query_parser: (bool * string criterion) list parser
52925
71e938856a03 more robust read_query;
wenzelm
parents: 52865
diff changeset
    19
  val read_query: Position.T -> string -> (bool * string criterion) list
30785
15f64e05e703 Limit the number of results returned by auto_solves.
Timothy Bourke
parents: 30693
diff changeset
    20
  val find_theorems: Proof.context -> thm option -> int option -> bool ->
43067
76e1d25c6357 separate query parsing from actual search
krauss
parents: 42669
diff changeset
    21
    (bool * term criterion) list -> int option * (Facts.ref * thm) list
76e1d25c6357 separate query parsing from actual search
krauss
parents: 42669
diff changeset
    22
  val find_theorems_cmd: Proof.context -> thm option -> int option -> bool ->
30785
15f64e05e703 Limit the number of results returned by auto_solves.
Timothy Bourke
parents: 30693
diff changeset
    23
    (bool * string criterion) list -> int option * (Facts.ref * thm) list
30186
1f836e949ac2 replaced archaic Display.pretty_fact by FindTheorems.pretty_thm, which observes the context properly (as did the former prt_fact already);
wenzelm
parents: 30143
diff changeset
    24
  val pretty_thm: Proof.context -> Facts.ref * thm -> Pretty.T
62848
e4140efe699e clarified bootstrap -- more uniform use of ML files;
wenzelm
parents: 61841
diff changeset
    25
  val pretty_theorems: Proof.state ->
e4140efe699e clarified bootstrap -- more uniform use of ML files;
wenzelm
parents: 61841
diff changeset
    26
    int option -> bool -> (bool * string criterion) list -> Pretty.T
e4140efe699e clarified bootstrap -- more uniform use of ML files;
wenzelm
parents: 61841
diff changeset
    27
  val proof_state: Toplevel.state -> Proof.state
16033
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
    28
end;
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
    29
33301
1fe9fc908ec3 less hermetic ML;
wenzelm
parents: 33290
diff changeset
    30
structure Find_Theorems: FIND_THEOREMS =
16033
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
    31
struct
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
    32
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
    33
(** search criteria **)
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
    34
16036
1da07ac33711 added read_criterion/pretty_criterion;
wenzelm
parents: 16033
diff changeset
    35
datatype 'term criterion =
46717
b09afce1e54f removed broken/unused introiff (cf. d452117ba564);
wenzelm
parents: 46716
diff changeset
    36
  Name of string | Intro | Elim | Dest | Solves | Simp of 'term | Pattern of 'term;
16036
1da07ac33711 added read_criterion/pretty_criterion;
wenzelm
parents: 16033
diff changeset
    37
1da07ac33711 added read_criterion/pretty_criterion;
wenzelm
parents: 16033
diff changeset
    38
fun read_criterion _ (Name name) = Name name
1da07ac33711 added read_criterion/pretty_criterion;
wenzelm
parents: 16033
diff changeset
    39
  | read_criterion _ Intro = Intro
1da07ac33711 added read_criterion/pretty_criterion;
wenzelm
parents: 16033
diff changeset
    40
  | read_criterion _ Elim = Elim
1da07ac33711 added read_criterion/pretty_criterion;
wenzelm
parents: 16033
diff changeset
    41
  | read_criterion _ Dest = Dest
29857
2cc976ed8a3c FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
kleing
parents: 29848
diff changeset
    42
  | read_criterion _ Solves = Solves
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42358
diff changeset
    43
  | read_criterion ctxt (Simp str) = Simp (Proof_Context.read_term_pattern ctxt str)
59097
4b05ce4c84b0 revert "better" handling of abbreviation from c61fe520602b
haftmann
parents: 59096
diff changeset
    44
  | read_criterion ctxt (Pattern str) = Pattern (Proof_Context.read_term_pattern ctxt str);
16033
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
    45
16036
1da07ac33711 added read_criterion/pretty_criterion;
wenzelm
parents: 16033
diff changeset
    46
fun pretty_criterion ctxt (b, c) =
1da07ac33711 added read_criterion/pretty_criterion;
wenzelm
parents: 16033
diff changeset
    47
  let
1da07ac33711 added read_criterion/pretty_criterion;
wenzelm
parents: 16033
diff changeset
    48
    fun prfx s = if b then s else "-" ^ s;
1da07ac33711 added read_criterion/pretty_criterion;
wenzelm
parents: 16033
diff changeset
    49
  in
1da07ac33711 added read_criterion/pretty_criterion;
wenzelm
parents: 16033
diff changeset
    50
    (case c of
1da07ac33711 added read_criterion/pretty_criterion;
wenzelm
parents: 16033
diff changeset
    51
      Name name => Pretty.str (prfx "name: " ^ quote name)
1da07ac33711 added read_criterion/pretty_criterion;
wenzelm
parents: 16033
diff changeset
    52
    | Intro => Pretty.str (prfx "intro")
1da07ac33711 added read_criterion/pretty_criterion;
wenzelm
parents: 16033
diff changeset
    53
    | Elim => Pretty.str (prfx "elim")
1da07ac33711 added read_criterion/pretty_criterion;
wenzelm
parents: 16033
diff changeset
    54
    | Dest => Pretty.str (prfx "dest")
29857
2cc976ed8a3c FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
kleing
parents: 29848
diff changeset
    55
    | Solves => Pretty.str (prfx "solves")
16088
f084ba24de29 cleaned up select_match
kleing
parents: 16077
diff changeset
    56
    | Simp pat => Pretty.block [Pretty.str (prfx "simp:"), Pretty.brk 1,
24920
2a45e400fdad generic Syntax.pretty/string_of operations;
wenzelm
parents: 24683
diff changeset
    57
        Pretty.quote (Syntax.pretty_term ctxt (Term.show_dummy_patterns pat))]
56914
f371418b9641 tuned message;
wenzelm
parents: 56912
diff changeset
    58
    | Pattern pat => Pretty.enclose (prfx "\"") "\""
24920
2a45e400fdad generic Syntax.pretty/string_of operations;
wenzelm
parents: 24683
diff changeset
    59
        [Syntax.pretty_term ctxt (Term.show_dummy_patterns pat)])
16036
1da07ac33711 added read_criterion/pretty_criterion;
wenzelm
parents: 16033
diff changeset
    60
  end;
16033
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
    61
30142
8d6145694bb5 moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
wenzelm
parents: 29882
diff changeset
    62
43620
43a195a0279b tuned layout;
wenzelm
parents: 43076
diff changeset
    63
43070
0318781be055 explicit type for search queries
krauss
parents: 43069
diff changeset
    64
(** queries **)
0318781be055 explicit type for search queries
krauss
parents: 43069
diff changeset
    65
0318781be055 explicit type for search queries
krauss
parents: 43069
diff changeset
    66
type 'term query = {
0318781be055 explicit type for search queries
krauss
parents: 43069
diff changeset
    67
  goal: thm option,
0318781be055 explicit type for search queries
krauss
parents: 43069
diff changeset
    68
  limit: int option,
0318781be055 explicit type for search queries
krauss
parents: 43069
diff changeset
    69
  rem_dups: bool,
0318781be055 explicit type for search queries
krauss
parents: 43069
diff changeset
    70
  criteria: (bool * 'term criterion) list
0318781be055 explicit type for search queries
krauss
parents: 43069
diff changeset
    71
};
0318781be055 explicit type for search queries
krauss
parents: 43069
diff changeset
    72
0318781be055 explicit type for search queries
krauss
parents: 43069
diff changeset
    73
fun map_criteria f {goal, limit, rem_dups, criteria} =
46718
wenzelm
parents: 46717
diff changeset
    74
  {goal = goal, limit = limit, rem_dups = rem_dups, criteria = f criteria};
43070
0318781be055 explicit type for search queries
krauss
parents: 43069
diff changeset
    75
43620
43a195a0279b tuned layout;
wenzelm
parents: 43076
diff changeset
    76
16033
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
    77
(** search criterion filters **)
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
    78
16895
df67fc190e06 Sort search results in order of relevance, where relevance =
kleing
parents: 16486
diff changeset
    79
(*generated filters are to be of the form
55671
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
    80
  input: (Facts.ref * thm)
53632
96808429b9ec more useful sorting of find_thms results
kleing
parents: 52982
diff changeset
    81
  output: (p:int, s:int, t:int) option, where
16895
df67fc190e06 Sort search results in order of relevance, where relevance =
kleing
parents: 16486
diff changeset
    82
    NONE indicates no match
17106
2bd6c20cdda1 use theory instead of obsolete Sign.sg;
wenzelm
parents: 16964
diff changeset
    83
    p is the primary sorting criterion
53632
96808429b9ec more useful sorting of find_thms results
kleing
parents: 52982
diff changeset
    84
      (eg. size of term)
96808429b9ec more useful sorting of find_thms results
kleing
parents: 52982
diff changeset
    85
    s is the secondary sorting criterion
16895
df67fc190e06 Sort search results in order of relevance, where relevance =
kleing
parents: 16486
diff changeset
    86
      (eg. number of assumptions in the theorem)
53632
96808429b9ec more useful sorting of find_thms results
kleing
parents: 52982
diff changeset
    87
    t is the tertiary sorting criterion
16895
df67fc190e06 Sort search results in order of relevance, where relevance =
kleing
parents: 16486
diff changeset
    88
      (eg. size of the substitution for intro, elim and dest)
df67fc190e06 Sort search results in order of relevance, where relevance =
kleing
parents: 16486
diff changeset
    89
  when applying a set of filters to a thm, fold results in:
53632
96808429b9ec more useful sorting of find_thms results
kleing
parents: 52982
diff changeset
    90
    (max p, max s, sum of all t)
16895
df67fc190e06 Sort search results in order of relevance, where relevance =
kleing
parents: 16486
diff changeset
    91
*)
df67fc190e06 Sort search results in order of relevance, where relevance =
kleing
parents: 16486
diff changeset
    92
16088
f084ba24de29 cleaned up select_match
kleing
parents: 16077
diff changeset
    93
f084ba24de29 cleaned up select_match
kleing
parents: 16077
diff changeset
    94
(* matching theorems *)
17106
2bd6c20cdda1 use theory instead of obsolete Sign.sg;
wenzelm
parents: 16964
diff changeset
    95
59970
e9f73d87d904 proper context for Object_Logic operations;
wenzelm
parents: 59936
diff changeset
    96
fun is_nontrivial ctxt = Term.is_Const o Term.head_of o Object_Logic.drop_judgment ctxt;
16088
f084ba24de29 cleaned up select_match
kleing
parents: 16077
diff changeset
    97
16964
6a25e42eaff5 Ordering is now: first by number of assumptions, second by the substitution size.
kleing
parents: 16895
diff changeset
    98
(*extract terms from term_src, refine them to the parts that concern us,
6a25e42eaff5 Ordering is now: first by number of assumptions, second by the substitution size.
kleing
parents: 16895
diff changeset
    99
  if po try match them against obj else vice versa.
6a25e42eaff5 Ordering is now: first by number of assumptions, second by the substitution size.
kleing
parents: 16895
diff changeset
   100
  trivial matches are ignored.
6a25e42eaff5 Ordering is now: first by number of assumptions, second by the substitution size.
kleing
parents: 16895
diff changeset
   101
  returns: smallest substitution size*)
46717
b09afce1e54f removed broken/unused introiff (cf. d452117ba564);
wenzelm
parents: 46716
diff changeset
   102
fun is_matching_thm (extract_terms, refine_term) ctxt po obj term_src =
16088
f084ba24de29 cleaned up select_match
kleing
parents: 16077
diff changeset
   103
  let
42360
da8817d01e7c modernized structure Proof_Context;
wenzelm
parents: 42358
diff changeset
   104
    val thy = Proof_Context.theory_of ctxt;
16088
f084ba24de29 cleaned up select_match
kleing
parents: 16077
diff changeset
   105
16486
1a12cdb6ee6b get_thm(s): Name;
wenzelm
parents: 16088
diff changeset
   106
    fun matches pat =
59970
e9f73d87d904 proper context for Object_Logic operations;
wenzelm
parents: 59936
diff changeset
   107
      is_nontrivial ctxt pat andalso
46717
b09afce1e54f removed broken/unused introiff (cf. d452117ba564);
wenzelm
parents: 46716
diff changeset
   108
      Pattern.matches thy (if po then (pat, obj) else (obj, pat));
16895
df67fc190e06 Sort search results in order of relevance, where relevance =
kleing
parents: 16486
diff changeset
   109
52940
wenzelm
parents: 52928
diff changeset
   110
    fun subst_size pat =
18184
43c4589a9a78 tuned Pattern.match/unify;
wenzelm
parents: 17972
diff changeset
   111
      let val (_, subst) =
43c4589a9a78 tuned Pattern.match/unify;
wenzelm
parents: 17972
diff changeset
   112
        Pattern.match thy (if po then (pat, obj) else (obj, pat)) (Vartab.empty, Vartab.empty)
17205
8994750ae33c refer to theory instead of low-level tsig;
wenzelm
parents: 17106
diff changeset
   113
      in Vartab.fold (fn (_, (_, t)) => fn n => size_of_term t + n) subst 0 end;
16088
f084ba24de29 cleaned up select_match
kleing
parents: 16077
diff changeset
   114
52941
28407b5f1c72 tuned signature;
wenzelm
parents: 52940
diff changeset
   115
    fun best_match [] = NONE
28407b5f1c72 tuned signature;
wenzelm
parents: 52940
diff changeset
   116
      | best_match xs = SOME (foldl1 Int.min xs);
16895
df67fc190e06 Sort search results in order of relevance, where relevance =
kleing
parents: 16486
diff changeset
   117
16964
6a25e42eaff5 Ordering is now: first by number of assumptions, second by the substitution size.
kleing
parents: 16895
diff changeset
   118
    val match_thm = matches o refine_term;
16486
1a12cdb6ee6b get_thm(s): Name;
wenzelm
parents: 16088
diff changeset
   119
  in
52940
wenzelm
parents: 52928
diff changeset
   120
    map (subst_size o refine_term) (filter match_thm (extract_terms term_src))
52941
28407b5f1c72 tuned signature;
wenzelm
parents: 52940
diff changeset
   121
    |> best_match
16088
f084ba24de29 cleaned up select_match
kleing
parents: 16077
diff changeset
   122
  end;
f084ba24de29 cleaned up select_match
kleing
parents: 16077
diff changeset
   123
f084ba24de29 cleaned up select_match
kleing
parents: 16077
diff changeset
   124
16033
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
   125
(* filter_name *)
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
   126
55671
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   127
fun filter_name str_pat (thmref, _) =
70545
b93ba98e627a tuned signature;
wenzelm
parents: 67721
diff changeset
   128
  if match_string str_pat (Facts.ref_name thmref)
53632
96808429b9ec more useful sorting of find_thms results
kleing
parents: 52982
diff changeset
   129
  then SOME (0, 0, 0) else NONE;
16033
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
   130
30142
8d6145694bb5 moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
wenzelm
parents: 29882
diff changeset
   131
29857
2cc976ed8a3c FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
kleing
parents: 29848
diff changeset
   132
(* filter intro/elim/dest/solves rules *)
16033
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
   133
55671
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   134
fun filter_dest ctxt goal (_, thm) =
16033
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
   135
  let
16964
6a25e42eaff5 Ordering is now: first by number of assumptions, second by the substitution size.
kleing
parents: 16895
diff changeset
   136
    val extract_dest =
55671
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   137
     (fn thm => if Thm.no_prems thm then [] else [Thm.full_prop_of thm],
16033
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
   138
      hd o Logic.strip_imp_prems);
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
   139
    val prems = Logic.prems_of_goal goal 1;
16895
df67fc190e06 Sort search results in order of relevance, where relevance =
kleing
parents: 16486
diff changeset
   140
55671
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   141
    fun try_subst prem = is_matching_thm extract_dest ctxt true prem thm;
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19476
diff changeset
   142
    val successful = prems |> map_filter try_subst;
16033
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
   143
  in
16895
df67fc190e06 Sort search results in order of relevance, where relevance =
kleing
parents: 16486
diff changeset
   144
    (*if possible, keep best substitution (one with smallest size)*)
17106
2bd6c20cdda1 use theory instead of obsolete Sign.sg;
wenzelm
parents: 16964
diff changeset
   145
    (*dest rules always have assumptions, so a dest with one
16895
df67fc190e06 Sort search results in order of relevance, where relevance =
kleing
parents: 16486
diff changeset
   146
      assumption is as good as an intro rule with none*)
55671
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   147
    if not (null successful) then
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   148
      SOME (size_of_term (Thm.prop_of thm), Thm.nprems_of thm - 1, foldl1 Int.min successful)
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   149
    else NONE
16033
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
   150
  end;
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
   151
55671
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   152
fun filter_intro ctxt goal (_, thm) =
16033
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
   153
  let
55671
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   154
    val extract_intro = (single o Thm.full_prop_of, Logic.strip_imp_concl);
16036
1da07ac33711 added read_criterion/pretty_criterion;
wenzelm
parents: 16033
diff changeset
   155
    val concl = Logic.concl_of_goal goal 1;
16033
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
   156
  in
55671
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   157
    (case is_matching_thm extract_intro ctxt true concl thm of
59096
15f7ebb29d38 tuned variable names
haftmann
parents: 59095
diff changeset
   158
      SOME k => SOME (size_of_term (Thm.prop_of thm), Thm.nprems_of thm, k)
55671
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   159
    | NONE => NONE)
16033
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
   160
  end;
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
   161
55671
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   162
fun filter_elim ctxt goal (_, thm) =
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   163
  if Thm.nprems_of thm > 0 then
16964
6a25e42eaff5 Ordering is now: first by number of assumptions, second by the substitution size.
kleing
parents: 16895
diff changeset
   164
    let
55671
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   165
      val rule = Thm.full_prop_of thm;
16964
6a25e42eaff5 Ordering is now: first by number of assumptions, second by the substitution size.
kleing
parents: 16895
diff changeset
   166
      val prems = Logic.prems_of_goal goal 1;
6a25e42eaff5 Ordering is now: first by number of assumptions, second by the substitution size.
kleing
parents: 16895
diff changeset
   167
      val goal_concl = Logic.concl_of_goal goal 1;
26283
e19a5a7e83f1 more precise Author line;
wenzelm
parents: 25992
diff changeset
   168
      val rule_mp = hd (Logic.strip_imp_prems rule);
16964
6a25e42eaff5 Ordering is now: first by number of assumptions, second by the substitution size.
kleing
parents: 16895
diff changeset
   169
      val rule_concl = Logic.strip_imp_concl rule;
57690
5b652fd305d4 tuned comment;
wenzelm
parents: 56914
diff changeset
   170
      fun combine t1 t2 = Const ("*combine*", dummyT --> dummyT) $ (t1 $ t2);  (* FIXME ?!? *)
16964
6a25e42eaff5 Ordering is now: first by number of assumptions, second by the substitution size.
kleing
parents: 16895
diff changeset
   171
      val rule_tree = combine rule_mp rule_concl;
26283
e19a5a7e83f1 more precise Author line;
wenzelm
parents: 25992
diff changeset
   172
      fun goal_tree prem = combine prem goal_concl;
46717
b09afce1e54f removed broken/unused introiff (cf. d452117ba564);
wenzelm
parents: 46716
diff changeset
   173
      fun try_subst prem = is_matching_thm (single, I) ctxt true (goal_tree prem) rule_tree;
19482
9f11af8f7ef9 tuned basic list operators (flat, maps, map_filter);
wenzelm
parents: 19476
diff changeset
   174
      val successful = prems |> map_filter try_subst;
16964
6a25e42eaff5 Ordering is now: first by number of assumptions, second by the substitution size.
kleing
parents: 16895
diff changeset
   175
    in
32798
4b85b59a9f66 misc tuning and simplification;
wenzelm
parents: 32738
diff changeset
   176
      (*elim rules always have assumptions, so an elim with one
4b85b59a9f66 misc tuning and simplification;
wenzelm
parents: 32738
diff changeset
   177
        assumption is as good as an intro rule with none*)
59970
e9f73d87d904 proper context for Object_Logic operations;
wenzelm
parents: 59936
diff changeset
   178
      if is_nontrivial ctxt (Thm.major_prem_of thm) andalso not (null successful) then
55671
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   179
        SOME (size_of_term (Thm.prop_of thm), Thm.nprems_of thm - 1, foldl1 Int.min successful)
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   180
      else NONE
16964
6a25e42eaff5 Ordering is now: first by number of assumptions, second by the substitution size.
kleing
parents: 16895
diff changeset
   181
    end
46718
wenzelm
parents: 46717
diff changeset
   182
  else NONE;
16036
1da07ac33711 added read_criterion/pretty_criterion;
wenzelm
parents: 16033
diff changeset
   183
30143
98a986b02022 observe basic Isabelle/ML coding conventions;
wenzelm
parents: 30142
diff changeset
   184
fun filter_solves ctxt goal =
98a986b02022 observe basic Isabelle/ML coding conventions;
wenzelm
parents: 30142
diff changeset
   185
  let
64983
481b2855ee9a suppress warnings uniformly, even in explicit commands 'find_theorems', 'solve_direct';
wenzelm
parents: 64556
diff changeset
   186
    val thy' = Proof_Context.theory_of ctxt
481b2855ee9a suppress warnings uniformly, even in explicit commands 'find_theorems', 'solve_direct';
wenzelm
parents: 64556
diff changeset
   187
      |> Context_Position.set_visible_global false;
481b2855ee9a suppress warnings uniformly, even in explicit commands 'find_theorems', 'solve_direct';
wenzelm
parents: 64556
diff changeset
   188
    val ctxt' = Proof_Context.transfer thy' ctxt
481b2855ee9a suppress warnings uniformly, even in explicit commands 'find_theorems', 'solve_direct';
wenzelm
parents: 64556
diff changeset
   189
      |> Context_Position.set_visible false;
52704
b824497c8e86 modify background theory where it is actually required (cf. 51dfdcd88e84);
wenzelm
parents: 52703
diff changeset
   190
    val goal' = Thm.transfer thy' goal;
b824497c8e86 modify background theory where it is actually required (cf. 51dfdcd88e84);
wenzelm
parents: 52703
diff changeset
   191
52941
28407b5f1c72 tuned signature;
wenzelm
parents: 52940
diff changeset
   192
    fun limited_etac thm i =
67147
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 64984
diff changeset
   193
      Seq.take (Options.default_int \<^system_option>\<open>find_theorems_tactic_limit\<close>) o
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59098
diff changeset
   194
        eresolve_tac ctxt' [thm] i;
30143
98a986b02022 observe basic Isabelle/ML coding conventions;
wenzelm
parents: 30142
diff changeset
   195
    fun try_thm thm =
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 59098
diff changeset
   196
      if Thm.no_prems thm then resolve_tac ctxt' [thm] 1 goal'
54742
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53633
diff changeset
   197
      else
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53633
diff changeset
   198
        (limited_etac thm THEN_ALL_NEW (Goal.norm_hhf_tac ctxt' THEN' Method.assm_tac ctxt'))
7a86358a3c0b proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents: 53633
diff changeset
   199
          1 goal';
29857
2cc976ed8a3c FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
kleing
parents: 29848
diff changeset
   200
  in
55671
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   201
    fn (_, thm) =>
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   202
      if is_some (Seq.pull (try_thm thm))
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   203
      then SOME (size_of_term (Thm.prop_of thm), Thm.nprems_of thm, 0)
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   204
      else NONE
29857
2cc976ed8a3c FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
kleing
parents: 29848
diff changeset
   205
  end;
16033
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
   206
30142
8d6145694bb5 moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
wenzelm
parents: 29882
diff changeset
   207
16074
9e569163ba8c renamed search criterion 'rewrite' to 'simp'
kleing
parents: 16071
diff changeset
   208
(* filter_simp *)
16033
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
   209
55671
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   210
fun filter_simp ctxt t (_, thm) =
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   211
  let
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   212
    val mksimps = Simplifier.mksimps ctxt;
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   213
    val extract_simp =
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   214
      (map Thm.full_prop_of o mksimps, #1 o Logic.dest_equals o Logic.strip_imp_concl);
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   215
  in
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   216
    (case is_matching_thm extract_simp ctxt false t thm of
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   217
      SOME ss => SOME (size_of_term (Thm.prop_of thm), Thm.nprems_of thm, ss)
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   218
    | NONE => NONE)
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   219
  end;
16033
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
   220
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
   221
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
   222
(* filter_pattern *)
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
   223
59098
b6ba3adb48e3 eta-expand all search patterns using schematic place holders
haftmann
parents: 59097
diff changeset
   224
fun expand_abs t =
b6ba3adb48e3 eta-expand all search patterns using schematic place holders
haftmann
parents: 59097
diff changeset
   225
  let
b6ba3adb48e3 eta-expand all search patterns using schematic place holders
haftmann
parents: 59097
diff changeset
   226
    val m = Term.maxidx_of_term t + 1;
b6ba3adb48e3 eta-expand all search patterns using schematic place holders
haftmann
parents: 59097
diff changeset
   227
    val vs = strip_abs_vars t;
b6ba3adb48e3 eta-expand all search patterns using schematic place holders
haftmann
parents: 59097
diff changeset
   228
    val ts = map_index (fn (k, (_, T)) => Var ((Name.aT, m + k), T)) vs;
b6ba3adb48e3 eta-expand all search patterns using schematic place holders
haftmann
parents: 59097
diff changeset
   229
  in betapplys (t, ts) end;
b6ba3adb48e3 eta-expand all search patterns using schematic place holders
haftmann
parents: 59097
diff changeset
   230
32798
4b85b59a9f66 misc tuning and simplification;
wenzelm
parents: 32738
diff changeset
   231
fun get_names t = Term.add_const_names t (Term.add_free_names t []);
28900
53fd5cc685b4 FindTheorems performance improvements (from Timothy Bourke)
kleing
parents: 28211
diff changeset
   232
59095
3100a7b1c092 turn application-specific Pattern.matches_subterm into an application-private function
haftmann
parents: 59083
diff changeset
   233
(* Does pat match a subterm of obj? *)
3100a7b1c092 turn application-specific Pattern.matches_subterm into an application-private function
haftmann
parents: 59083
diff changeset
   234
fun matches_subterm thy (pat, obj) =
3100a7b1c092 turn application-specific Pattern.matches_subterm into an application-private function
haftmann
parents: 59083
diff changeset
   235
  let
3100a7b1c092 turn application-specific Pattern.matches_subterm into an application-private function
haftmann
parents: 59083
diff changeset
   236
    fun msub bounds obj = Pattern.matches thy (pat, obj) orelse
3100a7b1c092 turn application-specific Pattern.matches_subterm into an application-private function
haftmann
parents: 59083
diff changeset
   237
      (case obj of
59096
15f7ebb29d38 tuned variable names
haftmann
parents: 59095
diff changeset
   238
        Abs (_, T, t) => msub (bounds + 1) (snd (Term.dest_abs (Name.bound bounds, T, t)))
59095
3100a7b1c092 turn application-specific Pattern.matches_subterm into an application-private function
haftmann
parents: 59083
diff changeset
   239
      | t $ u => msub bounds t orelse msub bounds u
3100a7b1c092 turn application-specific Pattern.matches_subterm into an application-private function
haftmann
parents: 59083
diff changeset
   240
      | _ => false)
3100a7b1c092 turn application-specific Pattern.matches_subterm into an application-private function
haftmann
parents: 59083
diff changeset
   241
  in msub 0 obj end;
3100a7b1c092 turn application-specific Pattern.matches_subterm into an application-private function
haftmann
parents: 59083
diff changeset
   242
52940
wenzelm
parents: 52928
diff changeset
   243
(*Including all constants and frees is only sound because matching
wenzelm
parents: 52928
diff changeset
   244
  uses higher-order patterns. If full matching were used, then
wenzelm
parents: 52928
diff changeset
   245
  constants that may be subject to beta-reduction after substitution
wenzelm
parents: 52928
diff changeset
   246
  of frees should not be included for LHS set because they could be
wenzelm
parents: 52928
diff changeset
   247
  thrown away by the substituted function.  E.g. for (?F 1 2) do not
67721
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 67147
diff changeset
   248
  include 1 or 2, if it were possible for ?F to be (\<lambda>x y. 3).  The
52940
wenzelm
parents: 52928
diff changeset
   249
  largest possible set should always be included on the RHS.*)
30143
98a986b02022 observe basic Isabelle/ML coding conventions;
wenzelm
parents: 30142
diff changeset
   250
98a986b02022 observe basic Isabelle/ML coding conventions;
wenzelm
parents: 30142
diff changeset
   251
fun filter_pattern ctxt pat =
98a986b02022 observe basic Isabelle/ML coding conventions;
wenzelm
parents: 30142
diff changeset
   252
  let
59098
b6ba3adb48e3 eta-expand all search patterns using schematic place holders
haftmann
parents: 59097
diff changeset
   253
    val pat' = (expand_abs o Envir.eta_contract) pat;
b6ba3adb48e3 eta-expand all search patterns using schematic place holders
haftmann
parents: 59097
diff changeset
   254
    val pat_consts = get_names pat';
55671
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   255
    fun check ((x, thm), NONE) = check ((x, thm), SOME (get_names (Thm.full_prop_of thm)))
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   256
      | check ((_, thm), c as SOME thm_consts) =
33038
8f9594c31de4 dropped redundant gen_ prefix
haftmann
parents: 33037
diff changeset
   257
         (if subset (op =) (pat_consts, thm_consts) andalso
59098
b6ba3adb48e3 eta-expand all search patterns using schematic place holders
haftmann
parents: 59097
diff changeset
   258
            matches_subterm (Proof_Context.theory_of ctxt) (pat', Thm.full_prop_of thm)
55671
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   259
          then SOME (size_of_term (Thm.prop_of thm), Thm.nprems_of thm, 0) else NONE, c);
28900
53fd5cc685b4 FindTheorems performance improvements (from Timothy Bourke)
kleing
parents: 28211
diff changeset
   260
  in check end;
16033
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
   261
30142
8d6145694bb5 moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
wenzelm
parents: 29882
diff changeset
   262
16033
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
   263
(* interpret criteria as filters *)
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
   264
16036
1da07ac33711 added read_criterion/pretty_criterion;
wenzelm
parents: 16033
diff changeset
   265
local
1da07ac33711 added read_criterion/pretty_criterion;
wenzelm
parents: 16033
diff changeset
   266
1da07ac33711 added read_criterion/pretty_criterion;
wenzelm
parents: 16033
diff changeset
   267
fun err_no_goal c =
1da07ac33711 added read_criterion/pretty_criterion;
wenzelm
parents: 16033
diff changeset
   268
  error ("Current goal required for " ^ c ^ " search criterion");
1da07ac33711 added read_criterion/pretty_criterion;
wenzelm
parents: 16033
diff changeset
   269
28900
53fd5cc685b4 FindTheorems performance improvements (from Timothy Bourke)
kleing
parents: 28211
diff changeset
   270
fun filter_crit _ _ (Name name) = apfst (filter_name name)
16036
1da07ac33711 added read_criterion/pretty_criterion;
wenzelm
parents: 16033
diff changeset
   271
  | filter_crit _ NONE Intro = err_no_goal "intro"
1da07ac33711 added read_criterion/pretty_criterion;
wenzelm
parents: 16033
diff changeset
   272
  | filter_crit _ NONE Elim = err_no_goal "elim"
1da07ac33711 added read_criterion/pretty_criterion;
wenzelm
parents: 16033
diff changeset
   273
  | filter_crit _ NONE Dest = err_no_goal "dest"
29857
2cc976ed8a3c FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
kleing
parents: 29848
diff changeset
   274
  | filter_crit _ NONE Solves = err_no_goal "solves"
52940
wenzelm
parents: 52928
diff changeset
   275
  | filter_crit ctxt (SOME goal) Intro = apfst (filter_intro ctxt (Thm.prop_of goal))
wenzelm
parents: 52928
diff changeset
   276
  | filter_crit ctxt (SOME goal) Elim = apfst (filter_elim ctxt (Thm.prop_of goal))
wenzelm
parents: 52928
diff changeset
   277
  | filter_crit ctxt (SOME goal) Dest = apfst (filter_dest ctxt (Thm.prop_of goal))
29857
2cc976ed8a3c FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
kleing
parents: 29848
diff changeset
   278
  | filter_crit ctxt (SOME goal) Solves = apfst (filter_solves ctxt goal)
28900
53fd5cc685b4 FindTheorems performance improvements (from Timothy Bourke)
kleing
parents: 28211
diff changeset
   279
  | filter_crit ctxt _ (Simp pat) = apfst (filter_simp ctxt pat)
16088
f084ba24de29 cleaned up select_match
kleing
parents: 16077
diff changeset
   280
  | filter_crit ctxt _ (Pattern pat) = filter_pattern ctxt pat;
16036
1da07ac33711 added read_criterion/pretty_criterion;
wenzelm
parents: 16033
diff changeset
   281
53632
96808429b9ec more useful sorting of find_thms results
kleing
parents: 52982
diff changeset
   282
fun opt_not x = if is_some x then NONE else SOME (0, 0, 0);
16895
df67fc190e06 Sort search results in order of relevance, where relevance =
kleing
parents: 16486
diff changeset
   283
53632
96808429b9ec more useful sorting of find_thms results
kleing
parents: 52982
diff changeset
   284
fun opt_add (SOME (a, c, x)) (SOME (b, d, y)) = SOME (Int.max (a,b), Int.max (c, d), x + y : int)
26283
e19a5a7e83f1 more precise Author line;
wenzelm
parents: 25992
diff changeset
   285
  | opt_add _ _ = NONE;
16895
df67fc190e06 Sort search results in order of relevance, where relevance =
kleing
parents: 16486
diff changeset
   286
30143
98a986b02022 observe basic Isabelle/ML coding conventions;
wenzelm
parents: 30142
diff changeset
   287
fun app_filters thm =
98a986b02022 observe basic Isabelle/ML coding conventions;
wenzelm
parents: 30142
diff changeset
   288
  let
28900
53fd5cc685b4 FindTheorems performance improvements (from Timothy Bourke)
kleing
parents: 28211
diff changeset
   289
    fun app (NONE, _, _) = NONE
32798
4b85b59a9f66 misc tuning and simplification;
wenzelm
parents: 32738
diff changeset
   290
      | app (SOME v, _, []) = SOME (v, thm)
30143
98a986b02022 observe basic Isabelle/ML coding conventions;
wenzelm
parents: 30142
diff changeset
   291
      | app (r, consts, f :: fs) =
98a986b02022 observe basic Isabelle/ML coding conventions;
wenzelm
parents: 30142
diff changeset
   292
          let val (r', consts') = f (thm, consts)
98a986b02022 observe basic Isabelle/ML coding conventions;
wenzelm
parents: 30142
diff changeset
   293
          in app (opt_add r r', consts', fs) end;
28900
53fd5cc685b4 FindTheorems performance improvements (from Timothy Bourke)
kleing
parents: 28211
diff changeset
   294
  in app end;
53fd5cc685b4 FindTheorems performance improvements (from Timothy Bourke)
kleing
parents: 28211
diff changeset
   295
16036
1da07ac33711 added read_criterion/pretty_criterion;
wenzelm
parents: 16033
diff changeset
   296
in
16033
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
   297
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
   298
fun filter_criterion ctxt opt_goal (b, c) =
28900
53fd5cc685b4 FindTheorems performance improvements (from Timothy Bourke)
kleing
parents: 28211
diff changeset
   299
  (if b then I else (apfst opt_not)) o filter_crit ctxt opt_goal c;
16895
df67fc190e06 Sort search results in order of relevance, where relevance =
kleing
parents: 16486
diff changeset
   300
55671
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   301
fun sorted_filter filters thms =
16895
df67fc190e06 Sort search results in order of relevance, where relevance =
kleing
parents: 16486
diff changeset
   302
  let
55671
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   303
    fun eval_filters thm = app_filters thm (SOME (0, 0, 0), NONE, filters);
16033
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
   304
53632
96808429b9ec more useful sorting of find_thms results
kleing
parents: 52982
diff changeset
   305
    (*filters return: (thm size, number of assumptions, substitution size) option, so
96808429b9ec more useful sorting of find_thms results
kleing
parents: 52982
diff changeset
   306
      sort according to size of thm first, then number of assumptions,
96808429b9ec more useful sorting of find_thms results
kleing
parents: 52982
diff changeset
   307
      then by the substitution size, then by term order *)
55671
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   308
    fun result_ord (((p0, s0, t0), (_, thm0)), ((p1, s1, t1), (_, thm1))) =
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   309
      prod_ord int_ord (prod_ord int_ord (prod_ord int_ord Term_Ord.term_ord))
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   310
         ((p1, (s1, (t1, Thm.full_prop_of thm1))), (p0, (s0, (t0, Thm.full_prop_of thm0))));
46977
bd0ee92cabe7 slightly more parallel find_theorems;
wenzelm
parents: 46961
diff changeset
   311
  in
55671
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   312
    grouped 100 Par_List.map eval_filters thms
46977
bd0ee92cabe7 slightly more parallel find_theorems;
wenzelm
parents: 46961
diff changeset
   313
    |> map_filter I |> sort result_ord |> map #2
bd0ee92cabe7 slightly more parallel find_theorems;
wenzelm
parents: 46961
diff changeset
   314
  end;
16033
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
   315
30822
8aac4b974280 superficial tuning;
wenzelm
parents: 30785
diff changeset
   316
fun lazy_filter filters =
8aac4b974280 superficial tuning;
wenzelm
parents: 30785
diff changeset
   317
  let
30785
15f64e05e703 Limit the number of results returned by auto_solves.
Timothy Bourke
parents: 30693
diff changeset
   318
    fun lazy_match thms = Seq.make (fn () => first_match thms)
15f64e05e703 Limit the number of results returned by auto_solves.
Timothy Bourke
parents: 30693
diff changeset
   319
    and first_match [] = NONE
30822
8aac4b974280 superficial tuning;
wenzelm
parents: 30785
diff changeset
   320
      | first_match (thm :: thms) =
53632
96808429b9ec more useful sorting of find_thms results
kleing
parents: 52982
diff changeset
   321
          (case app_filters thm (SOME (0, 0, 0), NONE, filters) of
30785
15f64e05e703 Limit the number of results returned by auto_solves.
Timothy Bourke
parents: 30693
diff changeset
   322
            NONE => first_match thms
30822
8aac4b974280 superficial tuning;
wenzelm
parents: 30785
diff changeset
   323
          | SOME (_, t) => SOME (t, lazy_match thms));
30785
15f64e05e703 Limit the number of results returned by auto_solves.
Timothy Bourke
parents: 30693
diff changeset
   324
  in lazy_match end;
30822
8aac4b974280 superficial tuning;
wenzelm
parents: 30785
diff changeset
   325
16036
1da07ac33711 added read_criterion/pretty_criterion;
wenzelm
parents: 16033
diff changeset
   326
end;
1da07ac33711 added read_criterion/pretty_criterion;
wenzelm
parents: 16033
diff changeset
   327
16033
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
   328
52940
wenzelm
parents: 52928
diff changeset
   329
(* removing duplicates, preferring nicer names, roughly O(n log n) *)
22340
275802767bf3 Remove duplicates from printed theorems in find_theorems
kleing
parents: 19502
diff changeset
   330
25226
502d8676cdd6 improved notion of 'nicer' fact names (observe some name space properties);
wenzelm
parents: 24920
diff changeset
   331
local
502d8676cdd6 improved notion of 'nicer' fact names (observe some name space properties);
wenzelm
parents: 24920
diff changeset
   332
27486
c61507a98bff prefer theorem names without numbers
huffman
parents: 27173
diff changeset
   333
val index_ord = option_ord (K EQUAL);
59916
wenzelm
parents: 59888
diff changeset
   334
val hidden_ord = bool_ord o apply2 Long_Name.is_hidden;
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58928
diff changeset
   335
val qual_ord = int_ord o apply2 Long_Name.qualification;
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58928
diff changeset
   336
val txt_ord = int_ord o apply2 size;
25226
502d8676cdd6 improved notion of 'nicer' fact names (observe some name space properties);
wenzelm
parents: 24920
diff changeset
   337
63080
8326aa594273 find dynamic facts as well, but static ones are preferred;
wenzelm
parents: 62969
diff changeset
   338
fun nicer_name ((a, x), i) ((b, y), j) =
8326aa594273 find dynamic facts as well, but static ones are preferred;
wenzelm
parents: 62969
diff changeset
   339
  (case bool_ord (a, b) of EQUAL =>
8326aa594273 find dynamic facts as well, but static ones are preferred;
wenzelm
parents: 62969
diff changeset
   340
    (case hidden_ord (x, y) of EQUAL =>
8326aa594273 find dynamic facts as well, but static ones are preferred;
wenzelm
parents: 62969
diff changeset
   341
      (case index_ord (i, j) of EQUAL =>
8326aa594273 find dynamic facts as well, but static ones are preferred;
wenzelm
parents: 62969
diff changeset
   342
        (case qual_ord (x, y) of EQUAL => txt_ord (x, y) | ord => ord)
8326aa594273 find dynamic facts as well, but static ones are preferred;
wenzelm
parents: 62969
diff changeset
   343
      | ord => ord)
27486
c61507a98bff prefer theorem names without numbers
huffman
parents: 27173
diff changeset
   344
    | ord => ord)
25226
502d8676cdd6 improved notion of 'nicer' fact names (observe some name space properties);
wenzelm
parents: 24920
diff changeset
   345
  | ord => ord) <> GREATER;
502d8676cdd6 improved notion of 'nicer' fact names (observe some name space properties);
wenzelm
parents: 24920
diff changeset
   346
29848
a7c164e228e1 Nicer names in FindTheorems.
Timothy Bourke
parents: 29302
diff changeset
   347
fun rem_cdups nicer xs =
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26283
diff changeset
   348
  let
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26283
diff changeset
   349
    fun rem_c rev_seen [] = rev rev_seen
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26283
diff changeset
   350
      | rem_c rev_seen [x] = rem_c (x :: rev_seen) []
55671
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   351
      | rem_c rev_seen ((x as ((n, thm), _)) :: (y as ((n', thm'), _)) :: rest) =
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   352
          if Thm.eq_thm_prop (thm, thm')
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   353
          then rem_c rev_seen ((if nicer n n' then x else y) :: rest)
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   354
          else rem_c (x :: rev_seen) (y :: rest);
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26283
diff changeset
   355
  in rem_c [] xs end;
25226
502d8676cdd6 improved notion of 'nicer' fact names (observe some name space properties);
wenzelm
parents: 24920
diff changeset
   356
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26283
diff changeset
   357
in
25226
502d8676cdd6 improved notion of 'nicer' fact names (observe some name space properties);
wenzelm
parents: 24920
diff changeset
   358
30143
98a986b02022 observe basic Isabelle/ML coding conventions;
wenzelm
parents: 30142
diff changeset
   359
fun nicer_shortest ctxt =
98a986b02022 observe basic Isabelle/ML coding conventions;
wenzelm
parents: 30142
diff changeset
   360
  let
56143
ed2b660a52a1 more accurate resolution of hybrid facts, which actually changes the sort order of results;
wenzelm
parents: 56141
diff changeset
   361
    fun extern_shortest name =
63080
8326aa594273 find dynamic facts as well, but static ones are preferred;
wenzelm
parents: 62969
diff changeset
   362
      let
8326aa594273 find dynamic facts as well, but static ones are preferred;
wenzelm
parents: 62969
diff changeset
   363
        val facts = Proof_Context.facts_of_fact ctxt name;
8326aa594273 find dynamic facts as well, but static ones are preferred;
wenzelm
parents: 62969
diff changeset
   364
        val space = Facts.space_of facts;
8326aa594273 find dynamic facts as well, but static ones are preferred;
wenzelm
parents: 62969
diff changeset
   365
      in (Facts.is_dynamic facts name, Name_Space.extern_shortest ctxt space name) end;
29848
a7c164e228e1 Nicer names in FindTheorems.
Timothy Bourke
parents: 29302
diff changeset
   366
a7c164e228e1 Nicer names in FindTheorems.
Timothy Bourke
parents: 29302
diff changeset
   367
    fun nicer (Facts.Named ((x, _), i)) (Facts.Named ((y, _), j)) =
55672
5e25cc741ab9 support for completion within the formal context;
wenzelm
parents: 55671
diff changeset
   368
          nicer_name (extern_shortest x, i) (extern_shortest y, j)
29848
a7c164e228e1 Nicer names in FindTheorems.
Timothy Bourke
parents: 29302
diff changeset
   369
      | nicer (Facts.Fact _) (Facts.Named _) = true
55670
95454b2980ee removed dead code;
wenzelm
parents: 55669
diff changeset
   370
      | nicer (Facts.Named _) (Facts.Fact _) = false
95454b2980ee removed dead code;
wenzelm
parents: 55669
diff changeset
   371
      | nicer (Facts.Fact _) (Facts.Fact _) = true;
29848
a7c164e228e1 Nicer names in FindTheorems.
Timothy Bourke
parents: 29302
diff changeset
   372
  in nicer end;
a7c164e228e1 Nicer names in FindTheorems.
Timothy Bourke
parents: 29302
diff changeset
   373
a7c164e228e1 Nicer names in FindTheorems.
Timothy Bourke
parents: 29302
diff changeset
   374
fun rem_thm_dups nicer xs =
52940
wenzelm
parents: 52928
diff changeset
   375
  (xs ~~ (1 upto length xs))
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58928
diff changeset
   376
  |> sort (Term_Ord.fast_term_ord o apply2 (Thm.full_prop_of o #2 o #1))
29848
a7c164e228e1 Nicer names in FindTheorems.
Timothy Bourke
parents: 29302
diff changeset
   377
  |> rem_cdups nicer
59058
a78612c67ec0 renamed "pairself" to "apply2", in accordance to @{apply 2};
wenzelm
parents: 58928
diff changeset
   378
  |> sort (int_ord o apply2 #2)
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26283
diff changeset
   379
  |> map #1;
22340
275802767bf3 Remove duplicates from printed theorems in find_theorems
kleing
parents: 19502
diff changeset
   380
26336
a0e2b706ce73 renamed datatype thmref to Facts.ref, tuned interfaces;
wenzelm
parents: 26283
diff changeset
   381
end;
22340
275802767bf3 Remove duplicates from printed theorems in find_theorems
kleing
parents: 19502
diff changeset
   382
275802767bf3 Remove duplicates from printed theorems in find_theorems
kleing
parents: 19502
diff changeset
   383
52941
28407b5f1c72 tuned signature;
wenzelm
parents: 52940
diff changeset
   384
28407b5f1c72 tuned signature;
wenzelm
parents: 52940
diff changeset
   385
(** main operations **)
28407b5f1c72 tuned signature;
wenzelm
parents: 52940
diff changeset
   386
28407b5f1c72 tuned signature;
wenzelm
parents: 52940
diff changeset
   387
(* filter_theorems *)
16033
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
   388
26283
e19a5a7e83f1 more precise Author line;
wenzelm
parents: 25992
diff changeset
   389
fun all_facts_of ctxt =
33381
81269c72321a find_theorems: respect conceal flag
krauss
parents: 33301
diff changeset
   390
  let
61054
add998b3c597 trim context for persistent storage;
wenzelm
parents: 60610
diff changeset
   391
    val thy = Proof_Context.theory_of ctxt;
add998b3c597 trim context for persistent storage;
wenzelm
parents: 60610
diff changeset
   392
    val transfer = Global_Theory.transfer_theories thy;
56158
c2c6d560e7b2 more explicit treatment of verbose mode, which includes concealed entries;
wenzelm
parents: 56143
diff changeset
   393
    val local_facts = Proof_Context.facts_of ctxt;
61054
add998b3c597 trim context for persistent storage;
wenzelm
parents: 60610
diff changeset
   394
    val global_facts = Global_Theory.facts_of thy;
56141
c06202417c4a back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
wenzelm
parents: 56140
diff changeset
   395
  in
63080
8326aa594273 find dynamic facts as well, but static ones are preferred;
wenzelm
parents: 62969
diff changeset
   396
   (Facts.dest_all (Context.Proof ctxt) false [global_facts] local_facts @
8326aa594273 find dynamic facts as well, but static ones are preferred;
wenzelm
parents: 62969
diff changeset
   397
     Facts.dest_all (Context.Proof ctxt) false [] global_facts)
61054
add998b3c597 trim context for persistent storage;
wenzelm
parents: 60610
diff changeset
   398
   |> maps Facts.selections
add998b3c597 trim context for persistent storage;
wenzelm
parents: 60610
diff changeset
   399
   |> map (apsnd transfer)
56141
c06202417c4a back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
wenzelm
parents: 56140
diff changeset
   400
  end;
17972
4969d6eb4c97 do not export find_thms;
wenzelm
parents: 17789
diff changeset
   401
43070
0318781be055 explicit type for search queries
krauss
parents: 43069
diff changeset
   402
fun filter_theorems ctxt theorems query =
16033
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
   403
  let
46718
wenzelm
parents: 46717
diff changeset
   404
    val {goal = opt_goal, limit = opt_limit, rem_dups, criteria} = query;
43069
88e45168272c moved questionable goal modification out of filter_theorems
krauss
parents: 43068
diff changeset
   405
    val filters = map (filter_criterion ctxt opt_goal) criteria;
16033
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
   406
41844
b933142e02d0 generalize find_theorems filters to work on raw propositions, too
krauss
parents: 41841
diff changeset
   407
    fun find_all theorems =
30785
15f64e05e703 Limit the number of results returned by auto_solves.
Timothy Bourke
parents: 30693
diff changeset
   408
      let
41844
b933142e02d0 generalize find_theorems filters to work on raw propositions, too
krauss
parents: 41841
diff changeset
   409
        val raw_matches = sorted_filter filters theorems;
30785
15f64e05e703 Limit the number of results returned by auto_solves.
Timothy Bourke
parents: 30693
diff changeset
   410
15f64e05e703 Limit the number of results returned by auto_solves.
Timothy Bourke
parents: 30693
diff changeset
   411
        val matches =
15f64e05e703 Limit the number of results returned by auto_solves.
Timothy Bourke
parents: 30693
diff changeset
   412
          if rem_dups
15f64e05e703 Limit the number of results returned by auto_solves.
Timothy Bourke
parents: 30693
diff changeset
   413
          then rem_thm_dups (nicer_shortest ctxt) raw_matches
15f64e05e703 Limit the number of results returned by auto_solves.
Timothy Bourke
parents: 30693
diff changeset
   414
          else raw_matches;
28900
53fd5cc685b4 FindTheorems performance improvements (from Timothy Bourke)
kleing
parents: 28211
diff changeset
   415
30785
15f64e05e703 Limit the number of results returned by auto_solves.
Timothy Bourke
parents: 30693
diff changeset
   416
        val len = length matches;
67147
dea94b1aabc3 prefer control symbol antiquotations;
wenzelm
parents: 64984
diff changeset
   417
        val lim = the_default (Options.default_int \<^system_option>\<open>find_theorems_limit\<close>) opt_limit;
34088
d6194ece49df avoid negative indices as argument ot drop
haftmann
parents: 33957
diff changeset
   418
      in (SOME len, drop (Int.max (len - lim, 0)) matches) end;
30785
15f64e05e703 Limit the number of results returned by auto_solves.
Timothy Bourke
parents: 30693
diff changeset
   419
15f64e05e703 Limit the number of results returned by auto_solves.
Timothy Bourke
parents: 30693
diff changeset
   420
    val find =
15f64e05e703 Limit the number of results returned by auto_solves.
Timothy Bourke
parents: 30693
diff changeset
   421
      if rem_dups orelse is_none opt_limit
15f64e05e703 Limit the number of results returned by auto_solves.
Timothy Bourke
parents: 30693
diff changeset
   422
      then find_all
30822
8aac4b974280 superficial tuning;
wenzelm
parents: 30785
diff changeset
   423
      else pair NONE o Seq.list_of o Seq.take (the opt_limit) o lazy_filter filters;
30785
15f64e05e703 Limit the number of results returned by auto_solves.
Timothy Bourke
parents: 30693
diff changeset
   424
41844
b933142e02d0 generalize find_theorems filters to work on raw propositions, too
krauss
parents: 41841
diff changeset
   425
  in find theorems end;
29857
2cc976ed8a3c FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
kleing
parents: 29848
diff changeset
   426
46718
wenzelm
parents: 46717
diff changeset
   427
fun filter_theorems_cmd ctxt theorems raw_query =
52941
28407b5f1c72 tuned signature;
wenzelm
parents: 52940
diff changeset
   428
  filter_theorems ctxt theorems (map_criteria (map (apsnd (read_criterion ctxt))) raw_query);
28407b5f1c72 tuned signature;
wenzelm
parents: 52940
diff changeset
   429
28407b5f1c72 tuned signature;
wenzelm
parents: 52940
diff changeset
   430
28407b5f1c72 tuned signature;
wenzelm
parents: 52940
diff changeset
   431
(* find_theorems *)
28407b5f1c72 tuned signature;
wenzelm
parents: 52940
diff changeset
   432
28407b5f1c72 tuned signature;
wenzelm
parents: 52940
diff changeset
   433
local
43067
76e1d25c6357 separate query parsing from actual search
krauss
parents: 42669
diff changeset
   434
76e1d25c6357 separate query parsing from actual search
krauss
parents: 42669
diff changeset
   435
fun gen_find_theorems filter ctxt opt_goal opt_limit rem_dups raw_criteria =
43069
88e45168272c moved questionable goal modification out of filter_theorems
krauss
parents: 43068
diff changeset
   436
  let
88e45168272c moved questionable goal modification out of filter_theorems
krauss
parents: 43068
diff changeset
   437
    val assms =
88e45168272c moved questionable goal modification out of filter_theorems
krauss
parents: 43068
diff changeset
   438
      Proof_Context.get_fact ctxt (Facts.named "local.assms")
88e45168272c moved questionable goal modification out of filter_theorems
krauss
parents: 43068
diff changeset
   439
        handle ERROR _ => [];
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61268
diff changeset
   440
    val add_prems = Seq.hd o TRY (Method.insert_tac ctxt assms 1);
43069
88e45168272c moved questionable goal modification out of filter_theorems
krauss
parents: 43068
diff changeset
   441
    val opt_goal' = Option.map add_prems opt_goal;
88e45168272c moved questionable goal modification out of filter_theorems
krauss
parents: 43068
diff changeset
   442
  in
55671
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   443
    filter ctxt (all_facts_of ctxt)
46718
wenzelm
parents: 46717
diff changeset
   444
      {goal = opt_goal', limit = opt_limit, rem_dups = rem_dups, criteria = raw_criteria}
43069
88e45168272c moved questionable goal modification out of filter_theorems
krauss
parents: 43068
diff changeset
   445
  end;
30186
1f836e949ac2 replaced archaic Display.pretty_fact by FindTheorems.pretty_thm, which observes the context properly (as did the former prt_fact already);
wenzelm
parents: 30143
diff changeset
   446
52941
28407b5f1c72 tuned signature;
wenzelm
parents: 52940
diff changeset
   447
in
28407b5f1c72 tuned signature;
wenzelm
parents: 52940
diff changeset
   448
43067
76e1d25c6357 separate query parsing from actual search
krauss
parents: 42669
diff changeset
   449
val find_theorems = gen_find_theorems filter_theorems;
76e1d25c6357 separate query parsing from actual search
krauss
parents: 42669
diff changeset
   450
val find_theorems_cmd = gen_find_theorems filter_theorems_cmd;
76e1d25c6357 separate query parsing from actual search
krauss
parents: 42669
diff changeset
   451
52941
28407b5f1c72 tuned signature;
wenzelm
parents: 52940
diff changeset
   452
end;
28407b5f1c72 tuned signature;
wenzelm
parents: 52940
diff changeset
   453
28407b5f1c72 tuned signature;
wenzelm
parents: 52940
diff changeset
   454
28407b5f1c72 tuned signature;
wenzelm
parents: 52940
diff changeset
   455
(* pretty_theorems *)
28407b5f1c72 tuned signature;
wenzelm
parents: 52940
diff changeset
   456
28407b5f1c72 tuned signature;
wenzelm
parents: 52940
diff changeset
   457
local
28407b5f1c72 tuned signature;
wenzelm
parents: 52940
diff changeset
   458
49888
ff2063be8227 more formal markup;
wenzelm
parents: 48646
diff changeset
   459
fun pretty_ref ctxt thmref =
ff2063be8227 more formal markup;
wenzelm
parents: 48646
diff changeset
   460
  let
ff2063be8227 more formal markup;
wenzelm
parents: 48646
diff changeset
   461
    val (name, sel) =
ff2063be8227 more formal markup;
wenzelm
parents: 48646
diff changeset
   462
      (case thmref of
ff2063be8227 more formal markup;
wenzelm
parents: 48646
diff changeset
   463
        Facts.Named ((name, _), sel) => (name, sel)
ff2063be8227 more formal markup;
wenzelm
parents: 48646
diff changeset
   464
      | Facts.Fact _ => raise Fail "Illegal literal fact");
ff2063be8227 more formal markup;
wenzelm
parents: 48646
diff changeset
   465
  in
63080
8326aa594273 find dynamic facts as well, but static ones are preferred;
wenzelm
parents: 62969
diff changeset
   466
    [Pretty.marks_str (#1 (Proof_Context.markup_extern_fact ctxt name), name),
56141
c06202417c4a back to a form of hybrid facts, to reduce performance impact of ed92ce2ac88e;
wenzelm
parents: 56140
diff changeset
   467
      Pretty.str (Facts.string_of_selection sel), Pretty.str ":", Pretty.brk 1]
49888
ff2063be8227 more formal markup;
wenzelm
parents: 48646
diff changeset
   468
  end;
ff2063be8227 more formal markup;
wenzelm
parents: 48646
diff changeset
   469
52941
28407b5f1c72 tuned signature;
wenzelm
parents: 52940
diff changeset
   470
in
28407b5f1c72 tuned signature;
wenzelm
parents: 52940
diff changeset
   471
55671
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   472
fun pretty_thm ctxt (thmref, thm) =
61268
abe08fb15a12 moved remaining display.ML to more_thm.ML;
wenzelm
parents: 61223
diff changeset
   473
  Pretty.block (pretty_ref ctxt thmref @ [Thm.pretty_thm ctxt thm]);
41845
6611b9cef38b reactivate time measurement (partly reverting c27b0b37041a);
krauss
parents: 41844
diff changeset
   474
52941
28407b5f1c72 tuned signature;
wenzelm
parents: 52940
diff changeset
   475
fun pretty_theorems state opt_limit rem_dups raw_criteria =
30143
98a986b02022 observe basic Isabelle/ML coding conventions;
wenzelm
parents: 30142
diff changeset
   476
  let
52941
28407b5f1c72 tuned signature;
wenzelm
parents: 52940
diff changeset
   477
    val ctxt = Proof.context_of state;
64984
wenzelm
parents: 64983
diff changeset
   478
    val opt_goal = try (#goal o Proof.simple_goal) state;
29857
2cc976ed8a3c FindTheorems: add solves feature; tidy up const name subsettin; patch by Timothy Bourke
kleing
parents: 29848
diff changeset
   479
    val criteria = map (apsnd (read_criterion ctxt)) raw_criteria;
52941
28407b5f1c72 tuned signature;
wenzelm
parents: 52940
diff changeset
   480
52940
wenzelm
parents: 52928
diff changeset
   481
    val (opt_found, theorems) =
55671
aeca05e62fef removed remains of old experiment (see b933142e02d0);
wenzelm
parents: 55670
diff changeset
   482
      filter_theorems ctxt (all_facts_of ctxt)
52855
fb1f026c48ff some actual find_theorems functionality;
wenzelm
parents: 52854
diff changeset
   483
        {goal = opt_goal, limit = opt_limit, rem_dups = rem_dups, criteria = criteria};
41845
6611b9cef38b reactivate time measurement (partly reverting c27b0b37041a);
krauss
parents: 41844
diff changeset
   484
    val returned = length theorems;
31684
d5d830979a54 minor tuning according to Isabelle/ML conventions;
wenzelm
parents: 31042
diff changeset
   485
30785
15f64e05e703 Limit the number of results returned by auto_solves.
Timothy Bourke
parents: 30693
diff changeset
   486
    val tally_msg =
52940
wenzelm
parents: 52928
diff changeset
   487
      (case opt_found of
38335
630f379f2660 misc tuning and simplification;
wenzelm
parents: 38334
diff changeset
   488
        NONE => "displaying " ^ string_of_int returned ^ " theorem(s)"
30822
8aac4b974280 superficial tuning;
wenzelm
parents: 30785
diff changeset
   489
      | SOME found =>
38335
630f379f2660 misc tuning and simplification;
wenzelm
parents: 38334
diff changeset
   490
          "found " ^ string_of_int found ^ " theorem(s)" ^
30822
8aac4b974280 superficial tuning;
wenzelm
parents: 30785
diff changeset
   491
            (if returned < found
8aac4b974280 superficial tuning;
wenzelm
parents: 30785
diff changeset
   492
             then " (" ^ string_of_int returned ^ " displayed)"
8aac4b974280 superficial tuning;
wenzelm
parents: 30785
diff changeset
   493
             else ""));
56912
293cd4dcfebc some position markup to help locating the query context, e.g. from "Info" dockable;
wenzelm
parents: 56908
diff changeset
   494
    val position_markup = Position.markup (Position.thread_data ()) Markup.position;
16033
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
   495
  in
56891
48899c43b07d tuned message -- more context for detached window etc.;
wenzelm
parents: 56621
diff changeset
   496
    Pretty.block
56912
293cd4dcfebc some position markup to help locating the query context, e.g. from "Info" dockable;
wenzelm
parents: 56908
diff changeset
   497
      (Pretty.fbreaks
293cd4dcfebc some position markup to help locating the query context, e.g. from "Info" dockable;
wenzelm
parents: 56908
diff changeset
   498
        (Pretty.mark position_markup (Pretty.keyword1 "find_theorems") ::
293cd4dcfebc some position markup to help locating the query context, e.g. from "Info" dockable;
wenzelm
parents: 56908
diff changeset
   499
          map (pretty_criterion ctxt) criteria)) ::
38335
630f379f2660 misc tuning and simplification;
wenzelm
parents: 38334
diff changeset
   500
    Pretty.str "" ::
56908
734f7e6151c9 tuned message: more compact, imitate actual command line;
wenzelm
parents: 56891
diff changeset
   501
    (if null theorems then [Pretty.str "found nothing"]
38335
630f379f2660 misc tuning and simplification;
wenzelm
parents: 38334
diff changeset
   502
     else
56908
734f7e6151c9 tuned message: more compact, imitate actual command line;
wenzelm
parents: 56891
diff changeset
   503
       Pretty.str (tally_msg ^ ":") ::
734f7e6151c9 tuned message: more compact, imitate actual command line;
wenzelm
parents: 56891
diff changeset
   504
       grouped 10 Par_List.map (Pretty.item o single o pretty_thm ctxt) (rev theorems))
52855
fb1f026c48ff some actual find_theorems functionality;
wenzelm
parents: 52854
diff changeset
   505
  end |> Pretty.fbreaks |> curry Pretty.blk 0;
30142
8d6145694bb5 moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
wenzelm
parents: 29882
diff changeset
   506
52941
28407b5f1c72 tuned signature;
wenzelm
parents: 52940
diff changeset
   507
end;
30142
8d6145694bb5 moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
wenzelm
parents: 29882
diff changeset
   508
32798
4b85b59a9f66 misc tuning and simplification;
wenzelm
parents: 32738
diff changeset
   509
46718
wenzelm
parents: 46717
diff changeset
   510
52865
02a7e7180ee5 slightly more general support for one-shot query operations via asynchronous print functions and temporary document overlay;
wenzelm
parents: 52863
diff changeset
   511
(** Isar command syntax **)
30142
8d6145694bb5 moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
wenzelm
parents: 29882
diff changeset
   512
8d6145694bb5 moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
wenzelm
parents: 29882
diff changeset
   513
local
8d6145694bb5 moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
wenzelm
parents: 29882
diff changeset
   514
8d6145694bb5 moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
wenzelm
parents: 29882
diff changeset
   515
val criterion =
62969
9f394a16c557 eliminated "xname" and variants;
wenzelm
parents: 62848
diff changeset
   516
  Parse.reserved "name" |-- Parse.!!! (Parse.$$$ ":" |-- Parse.name) >> Name ||
36950
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 35625
diff changeset
   517
  Parse.reserved "intro" >> K Intro ||
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 35625
diff changeset
   518
  Parse.reserved "elim" >> K Elim ||
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 35625
diff changeset
   519
  Parse.reserved "dest" >> K Dest ||
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 35625
diff changeset
   520
  Parse.reserved "solves" >> K Solves ||
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 35625
diff changeset
   521
  Parse.reserved "simp" |-- Parse.!!! (Parse.$$$ ":" |-- Parse.term) >> Simp ||
75b8f26f2f07 refer directly to structure Keyword and Parse;
wenzelm
parents: 35625
diff changeset
   522
  Parse.term >> Pattern;
30142
8d6145694bb5 moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
wenzelm
parents: 29882
diff changeset
   523
63429
baedd4724f08 tuned signature: more uniform Keyword.spec;
wenzelm
parents: 63080
diff changeset
   524
val query_keywords =
64556
851ae0e7b09c more symbols;
wenzelm
parents: 63429
diff changeset
   525
  Keyword.add_keywords [((":", \<^here>), Keyword.no_spec)] Keyword.empty_keywords;
58905
55160ef37e8f more frugal keywords;
wenzelm
parents: 58903
diff changeset
   526
30142
8d6145694bb5 moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
wenzelm
parents: 29882
diff changeset
   527
in
8d6145694bb5 moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
wenzelm
parents: 29882
diff changeset
   528
62848
e4140efe699e clarified bootstrap -- more uniform use of ML files;
wenzelm
parents: 61841
diff changeset
   529
val query_parser = Scan.repeat ((Scan.option Parse.minus >> is_none) -- criterion);
e4140efe699e clarified bootstrap -- more uniform use of ML files;
wenzelm
parents: 61841
diff changeset
   530
52925
71e938856a03 more robust read_query;
wenzelm
parents: 52865
diff changeset
   531
fun read_query pos str =
59083
88b0b1f28adc tuned signature;
wenzelm
parents: 59058
diff changeset
   532
  Token.explode query_keywords pos str
52855
fb1f026c48ff some actual find_theorems functionality;
wenzelm
parents: 52854
diff changeset
   533
  |> filter Token.is_proper
62848
e4140efe699e clarified bootstrap -- more uniform use of ML files;
wenzelm
parents: 61841
diff changeset
   534
  |> Scan.error (Scan.finite Token.stopper (Parse.!!! (query_parser --| Scan.ahead Parse.eof)))
52925
71e938856a03 more robust read_query;
wenzelm
parents: 52865
diff changeset
   535
  |> #1;
43068
ac769b5edd1d exported raw query parser; removed inconsistent clone
krauss
parents: 43067
diff changeset
   536
16033
f93ca3d4ffa7 Retrieve theorems from proof context -- improved version of
wenzelm
parents:
diff changeset
   537
end;
30142
8d6145694bb5 moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
wenzelm
parents: 29882
diff changeset
   538
52851
e71b5160f242 minimal print function "find_theorems", which merely echos its arguments;
wenzelm
parents: 52788
diff changeset
   539
e71b5160f242 minimal print function "find_theorems", which merely echos its arguments;
wenzelm
parents: 52788
diff changeset
   540
52865
02a7e7180ee5 slightly more general support for one-shot query operations via asynchronous print functions and temporary document overlay;
wenzelm
parents: 52863
diff changeset
   541
(** PIDE query operation **)
52854
92932931bd82 more general Output.result: allow to update arbitrary properties;
wenzelm
parents: 52851
diff changeset
   542
62848
e4140efe699e clarified bootstrap -- more uniform use of ML files;
wenzelm
parents: 61841
diff changeset
   543
fun proof_state st =
e4140efe699e clarified bootstrap -- more uniform use of ML files;
wenzelm
parents: 61841
diff changeset
   544
  (case try Toplevel.proof_of st of
e4140efe699e clarified bootstrap -- more uniform use of ML files;
wenzelm
parents: 61841
diff changeset
   545
    SOME state => state
e4140efe699e clarified bootstrap -- more uniform use of ML files;
wenzelm
parents: 61841
diff changeset
   546
  | NONE => Proof.init (Toplevel.context_of st));
e4140efe699e clarified bootstrap -- more uniform use of ML files;
wenzelm
parents: 61841
diff changeset
   547
52865
02a7e7180ee5 slightly more general support for one-shot query operations via asynchronous print functions and temporary document overlay;
wenzelm
parents: 52863
diff changeset
   548
val _ =
60610
f52b4b0c10c4 improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
wenzelm
parents: 59970
diff changeset
   549
  Query_Operation.register {name = "find_theorems", pri = Task_Queue.urgent_pri}
61223
dfccf6c06201 clarified markup;
wenzelm
parents: 61054
diff changeset
   550
    (fn {state = st, args, writeln_result, ...} =>
60610
f52b4b0c10c4 improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
wenzelm
parents: 59970
diff changeset
   551
      if can Toplevel.context_of st then
f52b4b0c10c4 improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
wenzelm
parents: 59970
diff changeset
   552
        let
f52b4b0c10c4 improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
wenzelm
parents: 59970
diff changeset
   553
          val [limit_arg, allow_dups_arg, query_arg] = args;
f52b4b0c10c4 improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
wenzelm
parents: 59970
diff changeset
   554
          val state = proof_state st;
f52b4b0c10c4 improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
wenzelm
parents: 59970
diff changeset
   555
          val opt_limit = Int.fromString limit_arg;
f52b4b0c10c4 improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
wenzelm
parents: 59970
diff changeset
   556
          val rem_dups = allow_dups_arg = "false";
f52b4b0c10c4 improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
wenzelm
parents: 59970
diff changeset
   557
          val criteria = read_query Position.none query_arg;
61223
dfccf6c06201 clarified markup;
wenzelm
parents: 61054
diff changeset
   558
        in writeln_result (Pretty.string_of (pretty_theorems state opt_limit rem_dups criteria)) end
60610
f52b4b0c10c4 improved scheduling for urgent tasks, using farm of replacement threads (may lead to factor 2 overloading, but CPUs are usually hyperthreaded);
wenzelm
parents: 59970
diff changeset
   559
      else error "Unknown context");
52851
e71b5160f242 minimal print function "find_theorems", which merely echos its arguments;
wenzelm
parents: 52788
diff changeset
   560
30142
8d6145694bb5 moved find_theorems.ML and find_consts.ML to Pure/Tools, collecting main implementation in one place each;
wenzelm
parents: 29882
diff changeset
   561
end;