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