src/Pure/Isar/find_theorems.ML
author wenzelm
Tue, 09 Oct 2007 00:20:13 +0200
changeset 24920 2a45e400fdad
parent 24683 c62320337a4e
child 25226 502d8676cdd6
permissions -rw-r--r--
generic Syntax.pretty/string_of operations;
wenzelm@16033
     1
(*  Title:      Pure/Isar/find_theorems.ML
wenzelm@16033
     2
    ID:         $Id$
wenzelm@16033
     3
    Author:     Rafal Kolanski, NICTA and Tobias Nipkow, TU Muenchen
wenzelm@16033
     4
wenzelm@16033
     5
Retrieve theorems from proof context.
wenzelm@16033
     6
*)
wenzelm@16033
     7
wenzelm@16033
     8
val thms_containing_limit = ref 40;
wenzelm@16033
     9
wenzelm@16033
    10
signature FIND_THEOREMS =
wenzelm@16033
    11
sig
wenzelm@16036
    12
  datatype 'term criterion =
kleing@16074
    13
    Name of string | Intro | Elim | Dest | Simp of 'term | Pattern of 'term
kleing@22340
    14
  val print_theorems: Proof.context -> term option -> int option -> bool ->
wenzelm@16036
    15
    (bool * string criterion) list -> unit
wenzelm@16033
    16
end;
wenzelm@16033
    17
wenzelm@16033
    18
structure FindTheorems: FIND_THEOREMS =
wenzelm@16033
    19
struct
wenzelm@16033
    20
wenzelm@16033
    21
(** search criteria **)
wenzelm@16033
    22
wenzelm@16036
    23
datatype 'term criterion =
kleing@16074
    24
  Name of string | Intro | Elim | Dest | Simp of 'term | Pattern of 'term;
wenzelm@16033
    25
wenzelm@16036
    26
fun read_criterion _ (Name name) = Name name
wenzelm@16036
    27
  | read_criterion _ Intro = Intro
wenzelm@16036
    28
  | read_criterion _ Elim = Elim
wenzelm@16036
    29
  | read_criterion _ Dest = Dest
wenzelm@24683
    30
  | read_criterion ctxt (Simp str) = Simp (ProofContext.read_term_pattern ctxt str)
wenzelm@24683
    31
  | read_criterion ctxt (Pattern str) = Pattern (ProofContext.read_term_pattern ctxt str);
wenzelm@16033
    32
wenzelm@16036
    33
fun pretty_criterion ctxt (b, c) =
wenzelm@16036
    34
  let
wenzelm@16036
    35
    fun prfx s = if b then s else "-" ^ s;
wenzelm@16036
    36
  in
wenzelm@16036
    37
    (case c of
wenzelm@16036
    38
      Name name => Pretty.str (prfx "name: " ^ quote name)
wenzelm@16036
    39
    | Intro => Pretty.str (prfx "intro")
wenzelm@16036
    40
    | Elim => Pretty.str (prfx "elim")
wenzelm@16036
    41
    | Dest => Pretty.str (prfx "dest")
kleing@16088
    42
    | Simp pat => Pretty.block [Pretty.str (prfx "simp:"), Pretty.brk 1,
wenzelm@24920
    43
        Pretty.quote (Syntax.pretty_term ctxt (Term.show_dummy_patterns pat))]
wenzelm@16036
    44
    | Pattern pat => Pretty.enclose (prfx " \"") "\""
wenzelm@24920
    45
        [Syntax.pretty_term ctxt (Term.show_dummy_patterns pat)])
wenzelm@16036
    46
  end;
wenzelm@16033
    47
wenzelm@16033
    48
wenzelm@16033
    49
wenzelm@16033
    50
(** search criterion filters **)
wenzelm@16033
    51
kleing@16895
    52
(*generated filters are to be of the form
wenzelm@17205
    53
  input: (thmref * thm)
wenzelm@17106
    54
  output: (p:int, s:int) option, where
kleing@16895
    55
    NONE indicates no match
wenzelm@17106
    56
    p is the primary sorting criterion
kleing@16895
    57
      (eg. number of assumptions in the theorem)
kleing@16895
    58
    s is the secondary sorting criterion
kleing@16895
    59
      (eg. size of the substitution for intro, elim and dest)
kleing@16895
    60
  when applying a set of filters to a thm, fold results in:
kleing@16895
    61
    (biggest p, sum of all s)
wenzelm@17106
    62
  currently p and s only matter for intro, elim, dest and simp filters,
wenzelm@17106
    63
  otherwise the default ordering is used.
kleing@16895
    64
*)
kleing@16895
    65
kleing@16088
    66
kleing@16088
    67
(* matching theorems *)
wenzelm@17106
    68
wenzelm@17205
    69
fun is_nontrivial thy = Term.is_Const o Term.head_of o ObjectLogic.drop_judgment thy;
kleing@16088
    70
kleing@16964
    71
(*extract terms from term_src, refine them to the parts that concern us,
kleing@16964
    72
  if po try match them against obj else vice versa.
kleing@16964
    73
  trivial matches are ignored.
kleing@16964
    74
  returns: smallest substitution size*)
kleing@16964
    75
fun is_matching_thm (extract_terms, refine_term) ctxt po obj term_src =
kleing@16088
    76
  let
wenzelm@17106
    77
    val thy = ProofContext.theory_of ctxt;
kleing@16088
    78
wenzelm@16486
    79
    fun matches pat =
wenzelm@17106
    80
      is_nontrivial thy pat andalso
wenzelm@17205
    81
      Pattern.matches thy (if po then (pat, obj) else (obj, pat));
kleing@16088
    82
kleing@16895
    83
    fun substsize pat =
wenzelm@18184
    84
      let val (_, subst) =
wenzelm@18184
    85
        Pattern.match thy (if po then (pat, obj) else (obj, pat)) (Vartab.empty, Vartab.empty)
wenzelm@17205
    86
      in Vartab.fold (fn (_, (_, t)) => fn n => size_of_term t + n) subst 0 end;
kleing@16895
    87
kleing@16895
    88
    fun bestmatch [] = NONE
wenzelm@17205
    89
     |  bestmatch xs = SOME (foldr1 Int.min xs);
kleing@16895
    90
kleing@16964
    91
    val match_thm = matches o refine_term;
wenzelm@16486
    92
  in
wenzelm@17106
    93
    map (substsize o refine_term)
wenzelm@19482
    94
        (filter match_thm (extract_terms term_src)) |> bestmatch
kleing@16088
    95
  end;
kleing@16088
    96
kleing@16088
    97
wenzelm@16033
    98
(* filter_name *)
wenzelm@16033
    99
wenzelm@17755
   100
fun match_string pat str =
wenzelm@17755
   101
  let
wenzelm@17755
   102
    fun match [] _ = true
wenzelm@17755
   103
      | match (p :: ps) s =
wenzelm@17755
   104
          size p <= size s andalso
wenzelm@17755
   105
            (case try (unprefix p) s of
wenzelm@17755
   106
              SOME s' => match ps s'
wenzelm@17755
   107
            | NONE => match (p :: ps) (String.substring (s, 1, size s - 1)));
wenzelm@17755
   108
  in match (space_explode "*" pat) str end;
wenzelm@16033
   109
wenzelm@17106
   110
fun filter_name str_pat (thmref, _) =
wenzelm@17755
   111
  if match_string str_pat (PureThy.name_of_thmref thmref)
wenzelm@17205
   112
  then SOME (0, 0) else NONE;
wenzelm@16033
   113
wenzelm@16033
   114
wenzelm@16036
   115
(* filter intro/elim/dest rules *)
wenzelm@16033
   116
wenzelm@17205
   117
fun filter_dest ctxt goal (_, thm) =
wenzelm@16033
   118
  let
kleing@16964
   119
    val extract_dest =
wenzelm@17205
   120
     (fn thm => if Thm.no_prems thm then [] else [Thm.full_prop_of thm],
wenzelm@16033
   121
      hd o Logic.strip_imp_prems);
wenzelm@16033
   122
    val prems = Logic.prems_of_goal goal 1;
kleing@16895
   123
kleing@16964
   124
    fun try_subst prem = is_matching_thm extract_dest ctxt true prem thm;
wenzelm@19482
   125
    val successful = prems |> map_filter try_subst;
wenzelm@16033
   126
  in
kleing@16895
   127
    (*if possible, keep best substitution (one with smallest size)*)
wenzelm@17106
   128
    (*dest rules always have assumptions, so a dest with one
kleing@16895
   129
      assumption is as good as an intro rule with none*)
wenzelm@17205
   130
    if not (null successful)
wenzelm@17205
   131
    then SOME (Thm.nprems_of thm - 1, foldr1 Int.min successful) else NONE
wenzelm@16033
   132
  end;
wenzelm@16033
   133
wenzelm@17205
   134
fun filter_intro ctxt goal (_, thm) =
wenzelm@16033
   135
  let
wenzelm@17205
   136
    val extract_intro = (single o Thm.full_prop_of, Logic.strip_imp_concl);
wenzelm@16036
   137
    val concl = Logic.concl_of_goal goal 1;
kleing@16964
   138
    val ss = is_matching_thm extract_intro ctxt true concl thm;
wenzelm@16033
   139
  in
wenzelm@18939
   140
    if is_some ss then SOME (Thm.nprems_of thm, the ss) else NONE
wenzelm@16033
   141
  end;
wenzelm@16033
   142
wenzelm@17205
   143
fun filter_elim ctxt goal (_, thm) =
kleing@16964
   144
  if not (Thm.no_prems thm) then
kleing@16964
   145
    let
wenzelm@17205
   146
      val rule = Thm.full_prop_of thm;
kleing@16964
   147
      val prems = Logic.prems_of_goal goal 1;
kleing@16964
   148
      val goal_concl = Logic.concl_of_goal goal 1;
kleing@16964
   149
      val rule_mp = (hd o Logic.strip_imp_prems) rule;
kleing@16964
   150
      val rule_concl = Logic.strip_imp_concl rule;
kleing@16964
   151
      fun combine t1 t2 = Const ("combine", dummyT --> dummyT) $ (t1 $ t2);
kleing@16964
   152
      val rule_tree = combine rule_mp rule_concl;
kleing@16964
   153
      fun goal_tree prem = (combine prem goal_concl);
wenzelm@17106
   154
      fun try_subst prem =
kleing@16964
   155
        is_matching_thm (single, I) ctxt true (goal_tree prem) rule_tree;
wenzelm@19482
   156
      val successful = prems |> map_filter try_subst;
kleing@16964
   157
    in
wenzelm@17106
   158
    (*elim rules always have assumptions, so an elim with one
kleing@16964
   159
      assumption is as good as an intro rule with none*)
wenzelm@17106
   160
      if is_nontrivial (ProofContext.theory_of ctxt) (Thm.major_prem_of thm)
wenzelm@17205
   161
        andalso not (null successful)
wenzelm@17205
   162
      then SOME (Thm.nprems_of thm - 1, foldr1 Int.min successful) else NONE
kleing@16964
   163
    end
kleing@16964
   164
  else NONE
wenzelm@16036
   165
wenzelm@16033
   166
kleing@16074
   167
(* filter_simp *)
wenzelm@16033
   168
wenzelm@17205
   169
fun filter_simp ctxt t (_, thm) =
wenzelm@16033
   170
  let
wenzelm@16033
   171
    val (_, {mk_rews = {mk, ...}, ...}) =
wenzelm@16033
   172
      MetaSimplifier.rep_ss (Simplifier.local_simpset_of ctxt);
wenzelm@17106
   173
    val extract_simp =
wenzelm@17205
   174
      (map Thm.full_prop_of o mk, #1 o Logic.dest_equals o Logic.strip_imp_concl);
kleing@16964
   175
    val ss = is_matching_thm extract_simp ctxt false t thm
wenzelm@17106
   176
  in
wenzelm@18939
   177
    if is_some ss then SOME (Thm.nprems_of thm, the ss) else NONE
kleing@16964
   178
  end;
wenzelm@16033
   179
wenzelm@16033
   180
wenzelm@16033
   181
(* filter_pattern *)
wenzelm@16033
   182
kleing@16088
   183
fun filter_pattern ctxt pat (_, thm) =
wenzelm@17205
   184
  if Pattern.matches_subterm (ProofContext.theory_of ctxt) (pat, Thm.full_prop_of thm)
wenzelm@17205
   185
  then SOME (0, 0) else NONE;
wenzelm@17205
   186
wenzelm@16033
   187
wenzelm@16033
   188
(* interpret criteria as filters *)
wenzelm@16033
   189
wenzelm@16036
   190
local
wenzelm@16036
   191
wenzelm@16036
   192
fun err_no_goal c =
wenzelm@16036
   193
  error ("Current goal required for " ^ c ^ " search criterion");
wenzelm@16036
   194
wenzelm@16033
   195
fun filter_crit _ _ (Name name) = filter_name name
wenzelm@16036
   196
  | filter_crit _ NONE Intro = err_no_goal "intro"
wenzelm@16036
   197
  | filter_crit _ NONE Elim = err_no_goal "elim"
wenzelm@16036
   198
  | filter_crit _ NONE Dest = err_no_goal "dest"
wenzelm@16036
   199
  | filter_crit ctxt (SOME goal) Intro = filter_intro ctxt goal
wenzelm@16036
   200
  | filter_crit ctxt (SOME goal) Elim = filter_elim ctxt goal
wenzelm@16036
   201
  | filter_crit ctxt (SOME goal) Dest = filter_dest ctxt goal
kleing@16088
   202
  | filter_crit ctxt _ (Simp pat) = filter_simp ctxt pat
kleing@16088
   203
  | filter_crit ctxt _ (Pattern pat) = filter_pattern ctxt pat;
wenzelm@16036
   204
wenzelm@19502
   205
fun opt_not x = if is_some x then NONE else SOME (0, 0);
kleing@16895
   206
wenzelm@17756
   207
fun opt_add (SOME (a, x)) (SOME (b, y)) = SOME (Int.max (a, b), x + y : int)
wenzelm@17205
   208
 |  opt_add _ _ = NONE;
kleing@16895
   209
wenzelm@16036
   210
in
wenzelm@16033
   211
wenzelm@16033
   212
fun filter_criterion ctxt opt_goal (b, c) =
kleing@16895
   213
  (if b then I else opt_not) o filter_crit ctxt opt_goal c;
wenzelm@16033
   214
kleing@16895
   215
fun all_filters filters thms =
kleing@16895
   216
  let
kleing@16895
   217
    fun eval_filters filters thm =
wenzelm@17205
   218
      fold opt_add (map (fn f => f thm) filters) (SOME (0, 0));
kleing@16895
   219
kleing@16895
   220
    (*filters return: (number of assumptions, substitution size) option, so
kleing@16964
   221
      sort (desc. in both cases) according to number of assumptions first,
kleing@16895
   222
      then by the substitution size*)
wenzelm@17205
   223
    fun thm_ord (((p0, s0), _), ((p1, s1), _)) =
wenzelm@17205
   224
      prod_ord int_ord int_ord ((p1, s1), (p0, s0));
kleing@16895
   225
  in
wenzelm@17205
   226
    map (`(eval_filters filters)) thms
wenzelm@19482
   227
    |> map_filter (fn (SOME x, y) => SOME (x, y) | (NONE, _) => NONE)
wenzelm@22360
   228
    |> sort thm_ord |> map #2
kleing@16895
   229
  end;
wenzelm@16033
   230
wenzelm@16036
   231
end;
wenzelm@16036
   232
wenzelm@16033
   233
kleing@22414
   234
(* removing duplicates, preferring nicer names, roughly n log n *)
kleing@22340
   235
kleing@22340
   236
fun rem_thm_dups xs =
wenzelm@22360
   237
  let
wenzelm@22360
   238
    fun nicer (Fact x) (Fact y) = size x <= size y
wenzelm@22360
   239
      | nicer (Fact _) _        = true
wenzelm@22360
   240
      | nicer (PureThy.Name x) (PureThy.Name y) = size x <= size y
wenzelm@22360
   241
      | nicer (PureThy.Name _) (Fact _) = false
wenzelm@22360
   242
      | nicer (PureThy.Name _) _ = true
wenzelm@22360
   243
      | nicer (NameSelection (x, _)) (NameSelection (y, _)) = size x <= size y
wenzelm@22360
   244
      | nicer (NameSelection _) _ = false;
kleing@22340
   245
kleing@22414
   246
    fun rem_cdups xs =
kleing@22414
   247
      let
kleing@22414
   248
        fun rem_c rev_seen [] = rev rev_seen
kleing@22414
   249
          | rem_c rev_seen [x] = rem_c (x::rev_seen) []
kleing@22414
   250
          | rem_c rev_seen ((x as ((n,t),_))::(y as ((n',t'),_))::xs) =
kleing@22414
   251
            if Thm.eq_thm_prop (t,t')
kleing@22414
   252
            then if nicer n n'
kleing@22414
   253
                 then rem_c rev_seen (x::xs)
kleing@22414
   254
                 else rem_c rev_seen (y::xs)
kleing@22414
   255
            else rem_c (x::rev_seen) (y::xs)
kleing@22414
   256
      in rem_c [] xs end;
kleing@22340
   257
kleing@22414
   258
  in ListPair.zip (xs, 1 upto length xs)
kleing@22414
   259
     |> sort (Term.fast_term_ord o pairself (prop_of o #2 o #1))
kleing@22414
   260
     |> rem_cdups
kleing@22414
   261
     |> sort (int_ord o pairself #2)
kleing@22414
   262
     |> map #1
kleing@22414
   263
  end;
kleing@22340
   264
kleing@22340
   265
wenzelm@16033
   266
(* print_theorems *)
wenzelm@16033
   267
wenzelm@17972
   268
fun find_thms ctxt spec =
wenzelm@18325
   269
  (PureThy.thms_containing (ProofContext.theory_of ctxt) spec
wenzelm@19482
   270
    |> maps PureThy.selections) @
wenzelm@18325
   271
  (ProofContext.lthms_containing ctxt spec
wenzelm@19482
   272
    |> maps PureThy.selections
wenzelm@19046
   273
    |> distinct (fn ((r1, th1), (r2, th2)) =>
wenzelm@22360
   274
        r1 = r2 andalso Thm.eq_thm_prop (th1, th2)));
wenzelm@17972
   275
kleing@22340
   276
fun print_theorems ctxt opt_goal opt_limit rem_dups raw_criteria =
wenzelm@16033
   277
  let
wenzelm@16036
   278
    val criteria = map (apsnd (read_criterion ctxt)) raw_criteria;
wenzelm@16036
   279
    val filters = map (filter_criterion ctxt opt_goal) criteria;
wenzelm@16036
   280
kleing@22340
   281
    val raw_matches = all_filters filters (find_thms ctxt ([], []));
wenzelm@22360
   282
    val matches =
kleing@22414
   283
      if rem_dups
wenzelm@22360
   284
      then rem_thm_dups raw_matches
wenzelm@22360
   285
      else raw_matches;
kleing@22340
   286
wenzelm@16036
   287
    val len = length matches;
wenzelm@19476
   288
    val limit = the_default (! thms_containing_limit) opt_limit;
kleing@22340
   289
    val thms = Library.drop (len - limit, matches);
wenzelm@16036
   290
wenzelm@16033
   291
    fun prt_fact (thmref, thm) =
wenzelm@16033
   292
      ProofContext.pretty_fact ctxt (PureThy.string_of_thmref thmref, [thm]);
wenzelm@16033
   293
  in
wenzelm@16036
   294
    Pretty.big_list "searched for:" (map (pretty_criterion ctxt) criteria) :: Pretty.str "" ::
kleing@22340
   295
     (if null thms then [Pretty.str "nothing found"]
wenzelm@16033
   296
      else
wenzelm@16036
   297
        [Pretty.str ("found " ^ string_of_int len ^ " theorems" ^
wenzelm@16036
   298
          (if len <= limit then "" else " (" ^ string_of_int limit ^ " displayed)") ^ ":"),
wenzelm@16036
   299
         Pretty.str ""] @
kleing@22340
   300
        map prt_fact thms)
wenzelm@16033
   301
    |> Pretty.chunks |> Pretty.writeln
wenzelm@16033
   302
  end;
wenzelm@16033
   303
wenzelm@16033
   304
end;