src/Pure/Isar/find_theorems.ML
author kleing
Fri, 13 Feb 2009 07:53:38 +1100
changeset 29819 29154e67731d
parent 29795 c8cee17d7e50
permissions -rw-r--r--
New command find_consts searching for constants by type (by Timothy Bourke).
wenzelm@16033
     1
(*  Title:      Pure/Isar/find_theorems.ML
wenzelm@26283
     2
    Author:     Rafal Kolanski and Gerwin Klein, NICTA
wenzelm@16033
     3
wenzelm@16033
     4
Retrieve theorems from proof context.
wenzelm@16033
     5
*)
wenzelm@16033
     6
wenzelm@16033
     7
signature FIND_THEOREMS =
wenzelm@16033
     8
sig
wenzelm@25992
     9
  val limit: int ref
kleing@29794
    10
  val tac_limit: int ref
kleing@29794
    11
wenzelm@16036
    12
  datatype 'term criterion =
kleing@29794
    13
    Name of string | Intro | Elim | Dest | Solves | Simp of 'term |
kleing@29794
    14
    Pattern of 'term
kleing@29794
    15
kleing@29794
    16
  val find_theorems: Proof.context -> thm option -> bool ->
kleing@29794
    17
    (bool * string criterion) list -> (Facts.ref * thm) list
kleing@29794
    18
kleing@29794
    19
  val print_theorems: Proof.context -> thm option -> int option -> bool ->
wenzelm@16036
    20
    (bool * string criterion) list -> unit
wenzelm@16033
    21
end;
wenzelm@16033
    22
wenzelm@16033
    23
structure FindTheorems: FIND_THEOREMS =
wenzelm@16033
    24
struct
wenzelm@16033
    25
wenzelm@16033
    26
(** search criteria **)
wenzelm@16033
    27
wenzelm@16036
    28
datatype 'term criterion =
kleing@29794
    29
  Name of string | Intro | Elim | Dest | Solves | Simp of 'term |
kleing@29794
    30
  Pattern of 'term;
wenzelm@16033
    31
wenzelm@16036
    32
fun read_criterion _ (Name name) = Name name
wenzelm@16036
    33
  | read_criterion _ Intro = Intro
wenzelm@16036
    34
  | read_criterion _ Elim = Elim
wenzelm@16036
    35
  | read_criterion _ Dest = Dest
kleing@29794
    36
  | read_criterion _ Solves = Solves
wenzelm@24683
    37
  | read_criterion ctxt (Simp str) = Simp (ProofContext.read_term_pattern ctxt str)
wenzelm@24683
    38
  | read_criterion ctxt (Pattern str) = Pattern (ProofContext.read_term_pattern ctxt str);
wenzelm@16033
    39
wenzelm@16036
    40
fun pretty_criterion ctxt (b, c) =
wenzelm@16036
    41
  let
wenzelm@16036
    42
    fun prfx s = if b then s else "-" ^ s;
wenzelm@16036
    43
  in
wenzelm@16036
    44
    (case c of
wenzelm@16036
    45
      Name name => Pretty.str (prfx "name: " ^ quote name)
wenzelm@16036
    46
    | Intro => Pretty.str (prfx "intro")
wenzelm@16036
    47
    | Elim => Pretty.str (prfx "elim")
wenzelm@16036
    48
    | Dest => Pretty.str (prfx "dest")
kleing@29794
    49
    | Solves => Pretty.str (prfx "solves")
kleing@16088
    50
    | Simp pat => Pretty.block [Pretty.str (prfx "simp:"), Pretty.brk 1,
wenzelm@24920
    51
        Pretty.quote (Syntax.pretty_term ctxt (Term.show_dummy_patterns pat))]
wenzelm@16036
    52
    | Pattern pat => Pretty.enclose (prfx " \"") "\""
wenzelm@24920
    53
        [Syntax.pretty_term ctxt (Term.show_dummy_patterns pat)])
wenzelm@16036
    54
  end;
wenzelm@16033
    55
wenzelm@16033
    56
(** search criterion filters **)
wenzelm@16033
    57
kleing@16895
    58
(*generated filters are to be of the form
wenzelm@26336
    59
  input: (Facts.ref * thm)
wenzelm@17106
    60
  output: (p:int, s:int) option, where
kleing@16895
    61
    NONE indicates no match
wenzelm@17106
    62
    p is the primary sorting criterion
kleing@16895
    63
      (eg. number of assumptions in the theorem)
kleing@16895
    64
    s is the secondary sorting criterion
kleing@16895
    65
      (eg. size of the substitution for intro, elim and dest)
kleing@16895
    66
  when applying a set of filters to a thm, fold results in:
kleing@16895
    67
    (biggest p, sum of all s)
wenzelm@17106
    68
  currently p and s only matter for intro, elim, dest and simp filters,
wenzelm@17106
    69
  otherwise the default ordering is used.
kleing@16895
    70
*)
kleing@16895
    71
kleing@16088
    72
kleing@16088
    73
(* matching theorems *)
wenzelm@17106
    74
wenzelm@17205
    75
fun is_nontrivial thy = Term.is_Const o Term.head_of o ObjectLogic.drop_judgment thy;
kleing@16088
    76
kleing@16964
    77
(*extract terms from term_src, refine them to the parts that concern us,
kleing@16964
    78
  if po try match them against obj else vice versa.
kleing@16964
    79
  trivial matches are ignored.
kleing@16964
    80
  returns: smallest substitution size*)
kleing@16964
    81
fun is_matching_thm (extract_terms, refine_term) ctxt po obj term_src =
kleing@16088
    82
  let
wenzelm@17106
    83
    val thy = ProofContext.theory_of ctxt;
kleing@16088
    84
wenzelm@16486
    85
    fun matches pat =
wenzelm@17106
    86
      is_nontrivial thy pat andalso
wenzelm@17205
    87
      Pattern.matches thy (if po then (pat, obj) else (obj, pat));
kleing@16088
    88
kleing@16895
    89
    fun substsize pat =
wenzelm@18184
    90
      let val (_, subst) =
wenzelm@18184
    91
        Pattern.match thy (if po then (pat, obj) else (obj, pat)) (Vartab.empty, Vartab.empty)
wenzelm@17205
    92
      in Vartab.fold (fn (_, (_, t)) => fn n => size_of_term t + n) subst 0 end;
kleing@16895
    93
kleing@16895
    94
    fun bestmatch [] = NONE
wenzelm@17205
    95
     |  bestmatch xs = SOME (foldr1 Int.min xs);
kleing@16895
    96
kleing@16964
    97
    val match_thm = matches o refine_term;
wenzelm@16486
    98
  in
wenzelm@26283
    99
    map (substsize o refine_term) (filter match_thm (extract_terms term_src))
wenzelm@26283
   100
    |> bestmatch
kleing@16088
   101
  end;
kleing@16088
   102
kleing@16088
   103
wenzelm@16033
   104
(* filter_name *)
wenzelm@16033
   105
wenzelm@17106
   106
fun filter_name str_pat (thmref, _) =
wenzelm@26336
   107
  if match_string str_pat (Facts.name_of_ref thmref)
wenzelm@17205
   108
  then SOME (0, 0) else NONE;
wenzelm@16033
   109
kleing@29794
   110
(* filter intro/elim/dest/solves rules *)
wenzelm@16033
   111
wenzelm@17205
   112
fun filter_dest ctxt goal (_, thm) =
wenzelm@16033
   113
  let
kleing@16964
   114
    val extract_dest =
wenzelm@17205
   115
     (fn thm => if Thm.no_prems thm then [] else [Thm.full_prop_of thm],
wenzelm@16033
   116
      hd o Logic.strip_imp_prems);
wenzelm@16033
   117
    val prems = Logic.prems_of_goal goal 1;
kleing@16895
   118
kleing@16964
   119
    fun try_subst prem = is_matching_thm extract_dest ctxt true prem thm;
wenzelm@19482
   120
    val successful = prems |> map_filter try_subst;
wenzelm@16033
   121
  in
kleing@16895
   122
    (*if possible, keep best substitution (one with smallest size)*)
wenzelm@17106
   123
    (*dest rules always have assumptions, so a dest with one
kleing@16895
   124
      assumption is as good as an intro rule with none*)
wenzelm@17205
   125
    if not (null successful)
wenzelm@17205
   126
    then SOME (Thm.nprems_of thm - 1, foldr1 Int.min successful) else NONE
wenzelm@16033
   127
  end;
wenzelm@16033
   128
wenzelm@17205
   129
fun filter_intro ctxt goal (_, thm) =
wenzelm@16033
   130
  let
wenzelm@17205
   131
    val extract_intro = (single o Thm.full_prop_of, Logic.strip_imp_concl);
wenzelm@16036
   132
    val concl = Logic.concl_of_goal goal 1;
kleing@16964
   133
    val ss = is_matching_thm extract_intro ctxt true concl thm;
wenzelm@16033
   134
  in
wenzelm@18939
   135
    if is_some ss then SOME (Thm.nprems_of thm, the ss) else NONE
wenzelm@16033
   136
  end;
wenzelm@16033
   137
wenzelm@17205
   138
fun filter_elim ctxt goal (_, thm) =
kleing@16964
   139
  if not (Thm.no_prems thm) then
kleing@16964
   140
    let
wenzelm@17205
   141
      val rule = Thm.full_prop_of thm;
kleing@16964
   142
      val prems = Logic.prems_of_goal goal 1;
kleing@16964
   143
      val goal_concl = Logic.concl_of_goal goal 1;
wenzelm@26283
   144
      val rule_mp = hd (Logic.strip_imp_prems rule);
kleing@16964
   145
      val rule_concl = Logic.strip_imp_concl rule;
wenzelm@26283
   146
      fun combine t1 t2 = Const ("*combine*", dummyT --> dummyT) $ (t1 $ t2);
kleing@16964
   147
      val rule_tree = combine rule_mp rule_concl;
wenzelm@26283
   148
      fun goal_tree prem = combine prem goal_concl;
wenzelm@17106
   149
      fun try_subst prem =
kleing@16964
   150
        is_matching_thm (single, I) ctxt true (goal_tree prem) rule_tree;
wenzelm@19482
   151
      val successful = prems |> map_filter try_subst;
kleing@16964
   152
    in
wenzelm@17106
   153
    (*elim rules always have assumptions, so an elim with one
kleing@16964
   154
      assumption is as good as an intro rule with none*)
wenzelm@17106
   155
      if is_nontrivial (ProofContext.theory_of ctxt) (Thm.major_prem_of thm)
wenzelm@17205
   156
        andalso not (null successful)
wenzelm@17205
   157
      then SOME (Thm.nprems_of thm - 1, foldr1 Int.min successful) else NONE
kleing@16964
   158
    end
kleing@16964
   159
  else NONE
wenzelm@16036
   160
kleing@29794
   161
val tac_limit = ref 5;
kleing@29794
   162
kleing@29794
   163
fun filter_solves ctxt goal = let
kleing@29794
   164
    val baregoal = Logic.get_goal (prop_of goal) 1;
kleing@29794
   165
kleing@29794
   166
    fun etacn thm i = Seq.take (!tac_limit) o etac thm i;
kleing@29794
   167
    fun try_thm thm = if Thm.no_prems thm then rtac thm 1 goal
kleing@29794
   168
                      else (etacn thm THEN_ALL_NEW
kleing@29794
   169
                             (Goal.norm_hhf_tac THEN'
kleing@29794
   170
                               Method.assumption_tac ctxt)) 1 goal;
kleing@29794
   171
  in
kleing@29794
   172
    fn (_, thm) => if (is_some o Seq.pull o try_thm) thm
kleing@29794
   173
                   then SOME (Thm.nprems_of thm, 0) else NONE
kleing@29794
   174
  end;
wenzelm@16033
   175
kleing@16074
   176
(* filter_simp *)
wenzelm@16033
   177
wenzelm@17205
   178
fun filter_simp ctxt t (_, thm) =
wenzelm@16033
   179
  let
wenzelm@16033
   180
    val (_, {mk_rews = {mk, ...}, ...}) =
wenzelm@29302
   181
      Simplifier.rep_ss (Simplifier.local_simpset_of ctxt);
wenzelm@17106
   182
    val extract_simp =
wenzelm@17205
   183
      (map Thm.full_prop_of o mk, #1 o Logic.dest_equals o Logic.strip_imp_concl);
kleing@16964
   184
    val ss = is_matching_thm extract_simp ctxt false t thm
wenzelm@17106
   185
  in
wenzelm@18939
   186
    if is_some ss then SOME (Thm.nprems_of thm, the ss) else NONE
kleing@16964
   187
  end;
wenzelm@16033
   188
wenzelm@16033
   189
wenzelm@16033
   190
(* filter_pattern *)
wenzelm@16033
   191
kleing@29794
   192
fun get_names t = (Term.add_const_names t []) union (Term.add_free_names t []);
kleing@29794
   193
fun get_thm_names (_, thm) = get_names (Thm.full_prop_of thm);
kleing@29794
   194
  (* Including all constants and frees is only sound because
kleing@29794
   195
     matching uses higher-order patterns. If full matching
kleing@29794
   196
     were used, then constants that may be subject to
kleing@29794
   197
     beta-reduction after substitution of frees should
kleing@29794
   198
     not be included for LHS set because they could be
kleing@29794
   199
     thrown away by the substituted function.
kleing@29794
   200
     e.g. for (?F 1 2) do not include 1 or 2, if it were
kleing@29794
   201
          possible for ?F to be (% x y. 3)
kleing@29794
   202
     The largest possible set should always be included on
kleing@29794
   203
     the RHS. *)
kleing@28900
   204
kleing@28900
   205
fun filter_pattern ctxt pat = let
kleing@29794
   206
    val pat_consts = get_names pat;
kleing@28900
   207
kleing@29794
   208
    fun check (t, NONE) = check (t, SOME (get_thm_names t))
kleing@28900
   209
      | check ((_, thm), c as SOME thm_consts) =
kleing@28900
   210
          (if pat_consts subset_string thm_consts
kleing@28900
   211
              andalso (Pattern.matches_subterm (ProofContext.theory_of ctxt)
kleing@28900
   212
                                               (pat, Thm.full_prop_of thm))
kleing@28900
   213
           then SOME (0, 0) else NONE, c);
kleing@28900
   214
  in check end;
wenzelm@16033
   215
wenzelm@16033
   216
(* interpret criteria as filters *)
wenzelm@16033
   217
wenzelm@16036
   218
local
wenzelm@16036
   219
wenzelm@16036
   220
fun err_no_goal c =
wenzelm@16036
   221
  error ("Current goal required for " ^ c ^ " search criterion");
wenzelm@16036
   222
kleing@29794
   223
val fix_goal = Thm.prop_of;
kleing@29794
   224
val fix_goalo = Option.map fix_goal;
kleing@29794
   225
kleing@28900
   226
fun filter_crit _ _ (Name name) = apfst (filter_name name)
wenzelm@16036
   227
  | filter_crit _ NONE Intro = err_no_goal "intro"
wenzelm@16036
   228
  | filter_crit _ NONE Elim = err_no_goal "elim"
wenzelm@16036
   229
  | filter_crit _ NONE Dest = err_no_goal "dest"
kleing@29794
   230
  | filter_crit _ NONE Solves = err_no_goal "solves"
kleing@29794
   231
  | filter_crit ctxt (SOME goal) Intro = apfst (filter_intro ctxt
kleing@29794
   232
                                                  (fix_goal goal))
kleing@29794
   233
  | filter_crit ctxt (SOME goal) Elim = apfst (filter_elim ctxt 
kleing@29794
   234
                                                  (fix_goal goal))
kleing@29794
   235
  | filter_crit ctxt (SOME goal) Dest = apfst (filter_dest ctxt
kleing@29794
   236
                                                  (fix_goal goal))
kleing@29794
   237
  | filter_crit ctxt (SOME goal) Solves = apfst (filter_solves ctxt goal)
kleing@28900
   238
  | filter_crit ctxt _ (Simp pat) = apfst (filter_simp ctxt pat)
kleing@16088
   239
  | filter_crit ctxt _ (Pattern pat) = filter_pattern ctxt pat;
wenzelm@16036
   240
wenzelm@19502
   241
fun opt_not x = if is_some x then NONE else SOME (0, 0);
kleing@16895
   242
wenzelm@17756
   243
fun opt_add (SOME (a, x)) (SOME (b, y)) = SOME (Int.max (a, b), x + y : int)
wenzelm@26283
   244
  | opt_add _ _ = NONE;
kleing@16895
   245
kleing@28900
   246
fun app_filters thm = let
kleing@28900
   247
    fun app (NONE, _, _) = NONE
kleing@28900
   248
      | app (SOME v, consts, []) = SOME (v, thm)
kleing@28900
   249
      | app (r, consts, f::fs) = let val (r', consts') = f (thm, consts)
kleing@28900
   250
                                 in app (opt_add r r', consts', fs) end;
kleing@28900
   251
  in app end;
kleing@28900
   252
wenzelm@16036
   253
in
wenzelm@16033
   254
wenzelm@16033
   255
fun filter_criterion ctxt opt_goal (b, c) =
kleing@28900
   256
  (if b then I else (apfst opt_not)) o filter_crit ctxt opt_goal c;
wenzelm@16033
   257
kleing@16895
   258
fun all_filters filters thms =
kleing@16895
   259
  let
kleing@28900
   260
    fun eval_filters thm = app_filters thm (SOME (0, 0), NONE, filters);
kleing@16895
   261
kleing@16895
   262
    (*filters return: (number of assumptions, substitution size) option, so
kleing@16964
   263
      sort (desc. in both cases) according to number of assumptions first,
kleing@16895
   264
      then by the substitution size*)
wenzelm@17205
   265
    fun thm_ord (((p0, s0), _), ((p1, s1), _)) =
wenzelm@17205
   266
      prod_ord int_ord int_ord ((p1, s1), (p0, s0));
kleing@28900
   267
  in map_filter eval_filters thms |> sort thm_ord |> map #2 end;
wenzelm@16033
   268
wenzelm@16036
   269
end;
wenzelm@16036
   270
wenzelm@16033
   271
kleing@22414
   272
(* removing duplicates, preferring nicer names, roughly n log n *)
kleing@22340
   273
wenzelm@25226
   274
local
wenzelm@25226
   275
huffman@27486
   276
val index_ord = option_ord (K EQUAL);
wenzelm@25226
   277
val hidden_ord = bool_ord o pairself NameSpace.is_hidden;
wenzelm@25226
   278
val qual_ord = int_ord o pairself (length o NameSpace.explode);
wenzelm@25226
   279
val txt_ord = int_ord o pairself size;
wenzelm@25226
   280
huffman@27486
   281
fun nicer_name (x, i) (y, j) =
huffman@27486
   282
  (case hidden_ord (x, y) of EQUAL =>
huffman@27486
   283
    (case index_ord (i, j) of EQUAL =>
huffman@27486
   284
      (case qual_ord (x, y) of EQUAL => txt_ord (x, y) | ord => ord)
huffman@27486
   285
    | ord => ord)
wenzelm@25226
   286
  | ord => ord) <> GREATER;
wenzelm@25226
   287
Timothy@29785
   288
fun rem_cdups nicer xs =
wenzelm@26336
   289
  let
wenzelm@26336
   290
    fun rem_c rev_seen [] = rev rev_seen
wenzelm@26336
   291
      | rem_c rev_seen [x] = rem_c (x :: rev_seen) []
wenzelm@26336
   292
      | rem_c rev_seen ((x as ((n, t), _)) :: (y as ((n', t'), _)) :: xs) =
wenzelm@26336
   293
        if Thm.eq_thm_prop (t, t')
wenzelm@26336
   294
        then rem_c rev_seen ((if nicer n n' then x else y) :: xs)
wenzelm@26336
   295
        else rem_c (x :: rev_seen) (y :: xs)
wenzelm@26336
   296
  in rem_c [] xs end;
wenzelm@26336
   297
wenzelm@25226
   298
in
wenzelm@25226
   299
Timothy@29785
   300
fun nicer_shortest ctxt = let
Timothy@29785
   301
    val ns = ProofContext.theory_of ctxt
Timothy@29785
   302
             |> PureThy.facts_of
Timothy@29785
   303
             |> Facts.space_of;
Timothy@29785
   304
Timothy@29785
   305
    val len_sort = sort (int_ord o (pairself size));
Timothy@29785
   306
    fun shorten s = (case len_sort (NameSpace.get_accesses ns s) of
Timothy@29785
   307
                       [] => s
Timothy@29785
   308
                     | s'::_ => s');
Timothy@29785
   309
Timothy@29785
   310
    fun nicer (Facts.Named ((x, _), i)) (Facts.Named ((y, _), j)) =
Timothy@29785
   311
          nicer_name (shorten x, i) (shorten y, j)
Timothy@29785
   312
      | nicer (Facts.Fact _) (Facts.Named _) = true
Timothy@29785
   313
      | nicer (Facts.Named _) (Facts.Fact _) = false;
Timothy@29785
   314
  in nicer end;
Timothy@29785
   315
Timothy@29785
   316
fun rem_thm_dups nicer xs =
wenzelm@26336
   317
  xs ~~ (1 upto length xs)
wenzelm@29269
   318
  |> sort (TermOrd.fast_term_ord o pairself (Thm.prop_of o #2 o #1))
Timothy@29785
   319
  |> rem_cdups nicer
wenzelm@26336
   320
  |> sort (int_ord o pairself #2)
wenzelm@26336
   321
  |> map #1;
wenzelm@25226
   322
wenzelm@25226
   323
end;
wenzelm@25226
   324
kleing@22340
   325
wenzelm@16033
   326
(* print_theorems *)
wenzelm@16033
   327
wenzelm@26283
   328
fun all_facts_of ctxt =
wenzelm@26336
   329
  maps Facts.selections
wenzelm@27173
   330
   (Facts.dest_static [] (PureThy.facts_of (ProofContext.theory_of ctxt)) @
wenzelm@27173
   331
    Facts.dest_static [] (ProofContext.facts_of ctxt));
wenzelm@17972
   332
wenzelm@25992
   333
val limit = ref 40;
wenzelm@25992
   334
kleing@29794
   335
fun find_theorems ctxt opt_goal rem_dups raw_criteria =
wenzelm@16033
   336
  let
kleing@29794
   337
    val add_prems = Seq.hd o (TRY (Method.insert_tac
kleing@29794
   338
                                     (Assumption.prems_of ctxt) 1));
kleing@29794
   339
    val opt_goal' = Option.map add_prems opt_goal;
kleing@29794
   340
wenzelm@16036
   341
    val criteria = map (apsnd (read_criterion ctxt)) raw_criteria;
kleing@29794
   342
    val filters = map (filter_criterion ctxt opt_goal') criteria;
wenzelm@16036
   343
wenzelm@26283
   344
    val raw_matches = all_filters filters (all_facts_of ctxt);
kleing@28900
   345
wenzelm@22360
   346
    val matches =
kleing@22414
   347
      if rem_dups
Timothy@29785
   348
      then rem_thm_dups (nicer_shortest ctxt) raw_matches
wenzelm@22360
   349
      else raw_matches;
kleing@29794
   350
  in matches end;
kleing@29794
   351
kleing@29794
   352
fun print_theorems ctxt opt_goal opt_limit rem_dups raw_criteria = let
kleing@29794
   353
    val start = start_timing ();
kleing@29794
   354
kleing@29794
   355
    val criteria = map (apsnd (read_criterion ctxt)) raw_criteria;
kleing@29794
   356
    val matches = find_theorems ctxt opt_goal rem_dups raw_criteria;
kleing@22340
   357
wenzelm@16036
   358
    val len = length matches;
wenzelm@25992
   359
    val lim = the_default (! limit) opt_limit;
wenzelm@25992
   360
    val thms = Library.drop (len - lim, matches);
wenzelm@16036
   361
kleing@28900
   362
    val end_msg = " in " ^
kleing@28900
   363
                  (List.nth (String.tokens Char.isSpace (end_timing start), 3))
kleing@28900
   364
                  ^ " secs"
wenzelm@16033
   365
  in
kleing@28900
   366
    Pretty.big_list "searched for:" (map (pretty_criterion ctxt) criteria)
kleing@29794
   367
        :: Pretty.str "" ::
kleing@28900
   368
     (if null thms then [Pretty.str ("nothing found" ^ end_msg)]
wenzelm@16033
   369
      else
wenzelm@16036
   370
        [Pretty.str ("found " ^ string_of_int len ^ " theorems" ^
kleing@28900
   371
          (if len <= lim then ""
kleing@28900
   372
           else " (" ^ string_of_int lim ^ " displayed)")
kleing@28900
   373
           ^ end_msg ^ ":"), Pretty.str ""] @
kleing@29795
   374
        map Display.pretty_fact thms)
wenzelm@16033
   375
    |> Pretty.chunks |> Pretty.writeln
kleing@29794
   376
  end
wenzelm@16033
   377
wenzelm@16033
   378
end;