| author | blanchet | 
| Wed, 08 Apr 2015 18:43:43 +0200 | |
| changeset 59959 | 1e3383a5204b | 
| parent 59917 | 9830c944670f | 
| child 61098 | e1b4b24f2ebd | 
| permissions | -rw-r--r-- | 
| 16014 | 1 | (* Title: Pure/simplifier.ML | 
| 2 | Author: Tobias Nipkow and Markus Wenzel, TU Muenchen | |
| 3 | ||
| 4 | Generic simplifier, suitable for most logics (see also | |
| 41386 | 5 | raw_simplifier.ML for the actual meta-level rewriting engine). | 
| 16014 | 6 | *) | 
| 7 | ||
| 8 | signature BASIC_SIMPLIFIER = | |
| 9 | sig | |
| 41228 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41226diff
changeset | 10 | include BASIC_RAW_SIMPLIFIER | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 11 | val simp_tac: Proof.context -> int -> tactic | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 12 | val asm_simp_tac: Proof.context -> int -> tactic | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 13 | val full_simp_tac: Proof.context -> int -> tactic | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 14 | val asm_lr_simp_tac: Proof.context -> int -> tactic | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 15 | val asm_full_simp_tac: Proof.context -> int -> tactic | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 16 | val safe_simp_tac: Proof.context -> int -> tactic | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 17 | val safe_asm_simp_tac: Proof.context -> int -> tactic | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 18 | val safe_full_simp_tac: Proof.context -> int -> tactic | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 19 | val safe_asm_lr_simp_tac: Proof.context -> int -> tactic | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 20 | val safe_asm_full_simp_tac: Proof.context -> int -> tactic | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 21 | val simplify: Proof.context -> thm -> thm | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 22 | val asm_simplify: Proof.context -> thm -> thm | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 23 | val full_simplify: Proof.context -> thm -> thm | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 24 | val asm_lr_simplify: Proof.context -> thm -> thm | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 25 | val asm_full_simplify: Proof.context -> thm -> thm | 
| 16014 | 26 | end; | 
| 27 | ||
| 28 | signature SIMPLIFIER = | |
| 29 | sig | |
| 30 | include BASIC_SIMPLIFIER | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 31 | val map_ss: (Proof.context -> Proof.context) -> Context.generic -> Context.generic | 
| 56510 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 32 | val attrib: (thm -> Proof.context -> Proof.context) -> attribute | 
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 33 | val simp_add: attribute | 
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 34 | val simp_del: attribute | 
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 35 | val cong_add: attribute | 
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 36 | val cong_del: attribute | 
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 37 | val check_simproc: Proof.context -> xstring * Position.T -> string | 
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 38 | val the_simproc: Proof.context -> string -> simproc | 
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 39 |   val def_simproc: {name: binding, lhss: term list,
 | 
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 40 | proc: morphism -> Proof.context -> cterm -> thm option, identifier: thm list} -> | 
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 41 | local_theory -> local_theory | 
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 42 |   val def_simproc_cmd: {name: binding, lhss: string list,
 | 
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 43 | proc: morphism -> Proof.context -> cterm -> thm option, identifier: thm list} -> | 
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 44 | local_theory -> local_theory | 
| 59917 
9830c944670f
more uniform "verbose" option to print name space;
 wenzelm parents: 
59621diff
changeset | 45 | val pretty_simpset: bool -> Proof.context -> Pretty.T | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 46 | val default_mk_sym: Proof.context -> thm -> thm option | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 47 | val prems_of: Proof.context -> thm list | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 48 | val add_simp: thm -> Proof.context -> Proof.context | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 49 | val del_simp: thm -> Proof.context -> Proof.context | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 50 | val add_eqcong: thm -> Proof.context -> Proof.context | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 51 | val del_eqcong: thm -> Proof.context -> Proof.context | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 52 | val add_cong: thm -> Proof.context -> Proof.context | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 53 | val del_cong: thm -> Proof.context -> Proof.context | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 54 | val add_prems: thm list -> Proof.context -> Proof.context | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 55 | val mksimps: Proof.context -> thm -> thm list | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 56 | val set_mksimps: (Proof.context -> thm -> thm list) -> Proof.context -> Proof.context | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 57 | val set_mkcong: (Proof.context -> thm -> thm) -> Proof.context -> Proof.context | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 58 | val set_mksym: (Proof.context -> thm -> thm option) -> Proof.context -> Proof.context | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 59 | val set_mkeqTrue: (Proof.context -> thm -> thm option) -> Proof.context -> Proof.context | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 60 | val set_termless: (term * term -> bool) -> Proof.context -> Proof.context | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 61 | val set_subgoaler: (Proof.context -> int -> tactic) -> Proof.context -> Proof.context | 
| 54729 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 wenzelm parents: 
54728diff
changeset | 62 | type trace_ops | 
| 54731 
384ac33802b0
clarified Trace_Ops: global theory data avoids init of simpset in Pure.thy, which is important to act as neutral element in merge;
 wenzelm parents: 
54729diff
changeset | 63 | val set_trace_ops: trace_ops -> theory -> theory | 
| 42795 
66fcc9882784
clarified map_simpset versus Simplifier.map_simpset_global;
 wenzelm parents: 
42793diff
changeset | 64 | val simproc_global_i: theory -> string -> term list -> | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 65 | (Proof.context -> term -> thm option) -> simproc | 
| 42795 
66fcc9882784
clarified map_simpset versus Simplifier.map_simpset_global;
 wenzelm parents: 
42793diff
changeset | 66 | val simproc_global: theory -> string -> string list -> | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 67 | (Proof.context -> term -> thm option) -> simproc | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 68 | val rewrite: Proof.context -> conv | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 69 | val asm_rewrite: Proof.context -> conv | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 70 | val full_rewrite: Proof.context -> conv | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 71 | val asm_lr_rewrite: Proof.context -> conv | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 72 | val asm_full_rewrite: Proof.context -> conv | 
| 30513 | 73 | val cong_modifiers: Method.modifier parser list | 
| 74 | val simp_modifiers': Method.modifier parser list | |
| 75 | val simp_modifiers: Method.modifier parser list | |
| 76 | val method_setup: Method.modifier parser list -> theory -> theory | |
| 18708 | 77 | val easy_setup: thm -> thm list -> theory -> theory | 
| 16014 | 78 | end; | 
| 79 | ||
| 80 | structure Simplifier: SIMPLIFIER = | |
| 81 | struct | |
| 82 | ||
| 41228 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41226diff
changeset | 83 | open Raw_Simplifier; | 
| 21708 | 84 | |
| 85 | ||
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 86 | (** declarations **) | 
| 16014 | 87 | |
| 26497 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 88 | (* attributes *) | 
| 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 89 | |
| 45620 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45375diff
changeset | 90 | fun attrib f = Thm.declaration_attribute (map_ss o f); | 
| 16014 | 91 | |
| 45620 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45375diff
changeset | 92 | val simp_add = attrib add_simp; | 
| 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45375diff
changeset | 93 | val simp_del = attrib del_simp; | 
| 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45375diff
changeset | 94 | val cong_add = attrib add_cong; | 
| 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45375diff
changeset | 95 | val cong_del = attrib del_cong; | 
| 26497 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 96 | |
| 16014 | 97 | |
| 22201 | 98 | (** named simprocs **) | 
| 99 | ||
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 100 | structure Simprocs = Generic_Data | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 101 | ( | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 102 | type T = simproc Name_Space.table; | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 103 | val empty : T = Name_Space.empty_table "simproc"; | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 104 | val extend = I; | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 105 | fun merge data : T = Name_Space.merge_tables data; | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 106 | ); | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 107 | |
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 108 | |
| 22204 | 109 | (* get simprocs *) | 
| 110 | ||
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 111 | val get_simprocs = Simprocs.get o Context.Proof; | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 112 | |
| 47005 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 wenzelm parents: 
47001diff
changeset | 113 | fun check_simproc ctxt = Name_Space.check (Context.Proof ctxt) (get_simprocs ctxt) #> #1; | 
| 42466 | 114 | val the_simproc = Name_Space.get o get_simprocs; | 
| 42465 
1ba52683512a
clarified check_simproc (with report) vs. the_simproc;
 wenzelm parents: 
42464diff
changeset | 115 | |
| 53171 | 116 | val _ = Theory.setup | 
| 56204 | 117 |   (ML_Antiquotation.value @{binding simproc}
 | 
| 53171 | 118 | (Args.context -- Scan.lift (Parse.position Args.name) | 
| 119 | >> (fn (ctxt, name) => | |
| 120 | "Simplifier.the_simproc ML_context " ^ ML_Syntax.print_string (check_simproc ctxt name)))); | |
| 22204 | 121 | |
| 122 | ||
| 123 | (* define simprocs *) | |
| 22201 | 124 | |
| 125 | local | |
| 126 | ||
| 42464 | 127 | fun gen_simproc prep {name = b, lhss, proc, identifier} lthy =
 | 
| 22201 | 128 | let | 
| 22236 | 129 | val simproc = make_simproc | 
| 47001 
a0e370d3d149
proper naming of simprocs according to actual target context;
 wenzelm parents: 
46776diff
changeset | 130 |       {name = Local_Theory.full_name lthy b,
 | 
| 22236 | 131 | lhss = | 
| 132 | let | |
| 133 | val lhss' = prep lthy lhss; | |
| 45326 
8fa859aebc0d
tuned -- Variable.declare_term is already part of Variable.auto_fixes;
 wenzelm parents: 
45291diff
changeset | 134 | val ctxt' = fold Variable.auto_fixes lhss' lthy; | 
| 59621 
291934bac95e
Thm.cterm_of and Thm.ctyp_of operate on local context;
 wenzelm parents: 
59573diff
changeset | 135 | in Variable.export_terms ctxt' lthy lhss' end |> map (Thm.cterm_of lthy), | 
| 22236 | 136 | proc = proc, | 
| 33551 
c40ced05b10a
define simprocs: do not apply target_morphism prematurely, this is already done in LocalTheory.declaration;
 wenzelm parents: 
33519diff
changeset | 137 | identifier = identifier}; | 
| 22201 | 138 | in | 
| 47001 
a0e370d3d149
proper naming of simprocs according to actual target context;
 wenzelm parents: 
46776diff
changeset | 139 |     lthy |> Local_Theory.declaration {syntax = false, pervasive = true} (fn phi => fn context =>
 | 
| 22201 | 140 | let | 
| 28991 | 141 | val b' = Morphism.binding phi b; | 
| 45290 | 142 | val simproc' = transform_simproc phi simproc; | 
| 22201 | 143 | in | 
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
42372diff
changeset | 144 | context | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 145 | |> Simprocs.map (#2 o Name_Space.define context true (b', simproc')) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 146 | |> map_ss (fn ctxt => ctxt addsimprocs [simproc']) | 
| 22201 | 147 | end) | 
| 148 | end; | |
| 149 | ||
| 150 | in | |
| 151 | ||
| 42464 | 152 | val def_simproc = gen_simproc Syntax.check_terms; | 
| 153 | val def_simproc_cmd = gen_simproc Syntax.read_terms; | |
| 22201 | 154 | |
| 155 | end; | |
| 156 | ||
| 157 | ||
| 158 | ||
| 56510 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 159 | (** pretty_simpset **) | 
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 160 | |
| 59917 
9830c944670f
more uniform "verbose" option to print name space;
 wenzelm parents: 
59621diff
changeset | 161 | fun pretty_simpset verbose ctxt = | 
| 56510 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 162 | let | 
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 163 | val pretty_term = Syntax.pretty_term ctxt; | 
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 164 | val pretty_thm = Display.pretty_thm ctxt; | 
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 165 | val pretty_thm_item = Display.pretty_thm_item ctxt; | 
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 166 | |
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 167 | fun pretty_simproc (name, lhss) = | 
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 168 | Pretty.block | 
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 169 | (Pretty.mark_str name :: Pretty.str ":" :: Pretty.fbrk :: | 
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 170 | Pretty.fbreaks (map (Pretty.item o single o pretty_term o Thm.term_of) lhss)); | 
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 171 | |
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 172 | fun pretty_cong_name (const, name) = | 
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 173 | pretty_term ((if const then Const else Free) (name, dummyT)); | 
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 174 | fun pretty_cong (name, thm) = | 
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 175 | Pretty.block [pretty_cong_name name, Pretty.str ":", Pretty.brk 1, pretty_thm thm]; | 
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 176 | |
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 177 |     val {simps, procs, congs, loopers, unsafe_solvers, safe_solvers, ...} =
 | 
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 178 | dest_ss (simpset_of ctxt); | 
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 179 | val simprocs = | 
| 59917 
9830c944670f
more uniform "verbose" option to print name space;
 wenzelm parents: 
59621diff
changeset | 180 | Name_Space.markup_entries verbose ctxt (Name_Space.space_of_table (get_simprocs ctxt)) procs; | 
| 56510 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 181 | in | 
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 182 | [Pretty.big_list "simplification rules:" (map (pretty_thm_item o #2) simps), | 
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 183 | Pretty.big_list "simplification procedures:" (map pretty_simproc simprocs), | 
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 184 | Pretty.big_list "congruences:" (map pretty_cong congs), | 
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 185 |       Pretty.strs ("loopers:" :: map quote loopers),
 | 
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 186 |       Pretty.strs ("unsafe solvers:" :: map quote unsafe_solvers),
 | 
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 187 |       Pretty.strs ("safe solvers:" :: map quote safe_solvers)]
 | 
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 188 | |> Pretty.chunks | 
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 189 | end; | 
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 190 | |
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 191 | |
| 
aec722524c33
added simproc markup, which also indicates legacy simprocs outside the name space;
 wenzelm parents: 
56204diff
changeset | 192 | |
| 17967 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 193 | (** simplification tactics and rules **) | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 194 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 195 | fun solve_all_tac solvers ctxt = | 
| 17967 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 196 | let | 
| 54728 | 197 |     val {subgoal_tac, ...} = Raw_Simplifier.internal_ss (simpset_of ctxt);
 | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 198 | val solve_tac = subgoal_tac (Raw_Simplifier.set_solvers solvers ctxt) THEN_ALL_NEW (K no_tac); | 
| 17967 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 199 | in DEPTH_SOLVE (solve_tac 1) end; | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 200 | |
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 201 | (*NOTE: may instantiate unknowns that appear also in other subgoals*) | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 202 | fun generic_simp_tac safe mode ctxt = | 
| 17967 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 203 | let | 
| 54728 | 204 |     val {loop_tacs, solvers = (unsafe_solvers, solvers), ...} =
 | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 205 | Raw_Simplifier.internal_ss (simpset_of ctxt); | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 206 | val loop_tac = FIRST' (map (fn (_, tac) => tac ctxt) (rev loop_tacs)); | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 207 | val solve_tac = FIRST' (map (Raw_Simplifier.solver ctxt) | 
| 22717 | 208 | (rev (if safe then solvers else unsafe_solvers))); | 
| 17967 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 209 | |
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 210 | fun simp_loop_tac i = | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 211 | Raw_Simplifier.generic_rewrite_goal_tac mode (solve_all_tac unsafe_solvers) ctxt i THEN | 
| 17967 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 212 | (solve_tac i ORELSE TRY ((loop_tac THEN_ALL_NEW simp_loop_tac) i)); | 
| 52458 
210bca64b894
less intrusive SELECT_GOAL: merely rearrange subgoals without detaching goal state, and thus preserve maxidx context;
 wenzelm parents: 
51717diff
changeset | 213 | in PREFER_GOAL (simp_loop_tac 1) end; | 
| 17967 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 214 | |
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 215 | local | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 216 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 217 | fun simp rew mode ctxt thm = | 
| 17967 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 218 | let | 
| 54728 | 219 |     val {solvers = (unsafe_solvers, _), ...} = Raw_Simplifier.internal_ss (simpset_of ctxt);
 | 
| 22717 | 220 | val tacf = solve_all_tac (rev unsafe_solvers); | 
| 17967 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 221 | fun prover s th = Option.map #1 (Seq.pull (tacf s th)); | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 222 | in rew mode prover ctxt thm end; | 
| 17967 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 223 | |
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 224 | in | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 225 | |
| 41228 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41226diff
changeset | 226 | val simp_thm = simp Raw_Simplifier.rewrite_thm; | 
| 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
41226diff
changeset | 227 | val simp_cterm = simp Raw_Simplifier.rewrite_cterm; | 
| 17967 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 228 | |
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 229 | end; | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 230 | |
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 231 | |
| 16806 | 232 | (* tactics *) | 
| 233 | ||
| 16014 | 234 | val simp_tac = generic_simp_tac false (false, false, false); | 
| 235 | val asm_simp_tac = generic_simp_tac false (false, true, false); | |
| 236 | val full_simp_tac = generic_simp_tac false (true, false, false); | |
| 237 | val asm_lr_simp_tac = generic_simp_tac false (true, true, false); | |
| 238 | val asm_full_simp_tac = generic_simp_tac false (true, true, true); | |
| 50107 | 239 | |
| 240 | (*not totally safe: may instantiate unknowns that appear also in other subgoals*) | |
| 241 | val safe_simp_tac = generic_simp_tac true (false, false, false); | |
| 242 | val safe_asm_simp_tac = generic_simp_tac true (false, true, false); | |
| 243 | val safe_full_simp_tac = generic_simp_tac true (true, false, false); | |
| 244 | val safe_asm_lr_simp_tac = generic_simp_tac true (true, true, false); | |
| 16014 | 245 | val safe_asm_full_simp_tac = generic_simp_tac true (true, true, true); | 
| 246 | ||
| 16806 | 247 | |
| 248 | (* conversions *) | |
| 249 | ||
| 17967 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 250 | val simplify = simp_thm (false, false, false); | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 251 | val asm_simplify = simp_thm (false, true, false); | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 252 | val full_simplify = simp_thm (true, false, false); | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 253 | val asm_lr_simplify = simp_thm (true, true, false); | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 254 | val asm_full_simplify = simp_thm (true, true, true); | 
| 16014 | 255 | |
| 17967 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 256 | val rewrite = simp_cterm (false, false, false); | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 257 | val asm_rewrite = simp_cterm (false, true, false); | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 258 | val full_rewrite = simp_cterm (true, false, false); | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 259 | val asm_lr_rewrite = simp_cterm (true, true, false); | 
| 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 wenzelm parents: 
17898diff
changeset | 260 | val asm_full_rewrite = simp_cterm (true, true, true); | 
| 16014 | 261 | |
| 262 | ||
| 263 | ||
| 264 | (** concrete syntax of attributes **) | |
| 265 | ||
| 266 | (* add / del *) | |
| 267 | ||
| 268 | val simpN = "simp"; | |
| 269 | val congN = "cong"; | |
| 270 | val onlyN = "only"; | |
| 271 | val no_asmN = "no_asm"; | |
| 272 | val no_asm_useN = "no_asm_use"; | |
| 273 | val no_asm_simpN = "no_asm_simp"; | |
| 274 | val asm_lrN = "asm_lr"; | |
| 275 | ||
| 276 | ||
| 24024 | 277 | (* simprocs *) | 
| 278 | ||
| 279 | local | |
| 280 | ||
| 281 | val add_del = | |
| 282 | (Args.del -- Args.colon >> K (op delsimprocs) || | |
| 283 | Scan.option (Args.add -- Args.colon) >> K (op addsimprocs)) | |
| 284 | >> (fn f => fn simproc => fn phi => Thm.declaration_attribute | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 285 | (K (Raw_Simplifier.map_ss (fn ctxt => f (ctxt, [transform_simproc phi simproc]))))); | 
| 24024 | 286 | |
| 287 | in | |
| 288 | ||
| 30528 | 289 | val simproc_att = | 
| 42465 
1ba52683512a
clarified check_simproc (with report) vs. the_simproc;
 wenzelm parents: 
42464diff
changeset | 290 | (Args.context -- Scan.lift add_del) :|-- (fn (ctxt, decl) => | 
| 
1ba52683512a
clarified check_simproc (with report) vs. the_simproc;
 wenzelm parents: 
42464diff
changeset | 291 | Scan.repeat1 (Scan.lift (Args.named_attribute (decl o the_simproc ctxt o check_simproc ctxt)))) | 
| 45375 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 wenzelm parents: 
45326diff
changeset | 292 | >> (fn atts => Thm.declaration_attribute (fn th => | 
| 46776 | 293 | fold (fn att => Thm.attribute_declaration (Morphism.form att) th) atts)); | 
| 24024 | 294 | |
| 295 | end; | |
| 24124 
4399175e3014
turned simp_depth_limit into configuration option;
 wenzelm parents: 
24024diff
changeset | 296 | |
| 24024 | 297 | |
| 16014 | 298 | (* conversions *) | 
| 299 | ||
| 300 | local | |
| 301 | ||
| 302 | fun conv_mode x = | |
| 303 | ((Args.parens (Args.$$$ no_asmN) >> K simplify || | |
| 304 | Args.parens (Args.$$$ no_asm_simpN) >> K asm_simplify || | |
| 305 | Args.parens (Args.$$$ no_asm_useN) >> K full_simplify || | |
| 306 | Scan.succeed asm_full_simplify) |> Scan.lift) x; | |
| 307 | ||
| 308 | in | |
| 309 | ||
| 30528 | 310 | val simplified = conv_mode -- Attrib.thms >> | 
| 311 | (fn (f, ths) => Thm.rule_attribute (fn context => | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 312 | f ((if null ths then I else Raw_Simplifier.clear_simpset) | 
| 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 313 | (Context.proof_of context) addsimps ths))); | 
| 16014 | 314 | |
| 315 | end; | |
| 316 | ||
| 317 | ||
| 318 | (* setup attributes *) | |
| 319 | ||
| 53171 | 320 | val _ = Theory.setup | 
| 56204 | 321 |  (Attrib.setup @{binding simp} (Attrib.add_del simp_add simp_del)
 | 
| 30528 | 322 | "declaration of Simplifier rewrite rule" #> | 
| 56204 | 323 |   Attrib.setup @{binding cong} (Attrib.add_del cong_add cong_del)
 | 
| 30528 | 324 | "declaration of Simplifier congruence rule" #> | 
| 56204 | 325 |   Attrib.setup @{binding simproc} simproc_att
 | 
| 33671 | 326 | "declaration of simplification procedures" #> | 
| 56204 | 327 |   Attrib.setup @{binding simplified} simplified "simplified rule");
 | 
| 16014 | 328 | |
| 329 | ||
| 330 | ||
| 31300 | 331 | (** method syntax **) | 
| 16014 | 332 | |
| 333 | val cong_modifiers = | |
| 58048 
aa6296d09e0e
more explicit Method.modifier with reported position;
 wenzelm parents: 
58008diff
changeset | 334 |  [Args.$$$ congN -- Args.colon >> K (Method.modifier cong_add @{here}),
 | 
| 
aa6296d09e0e
more explicit Method.modifier with reported position;
 wenzelm parents: 
58008diff
changeset | 335 |   Args.$$$ congN -- Args.add -- Args.colon >> K (Method.modifier cong_add @{here}),
 | 
| 
aa6296d09e0e
more explicit Method.modifier with reported position;
 wenzelm parents: 
58008diff
changeset | 336 |   Args.$$$ congN -- Args.del -- Args.colon >> K (Method.modifier cong_del @{here})];
 | 
| 16014 | 337 | |
| 338 | val simp_modifiers = | |
| 58048 
aa6296d09e0e
more explicit Method.modifier with reported position;
 wenzelm parents: 
58008diff
changeset | 339 |  [Args.$$$ simpN -- Args.colon >> K (Method.modifier simp_add @{here}),
 | 
| 
aa6296d09e0e
more explicit Method.modifier with reported position;
 wenzelm parents: 
58008diff
changeset | 340 |   Args.$$$ simpN -- Args.add -- Args.colon >> K (Method.modifier simp_add @{here}),
 | 
| 
aa6296d09e0e
more explicit Method.modifier with reported position;
 wenzelm parents: 
58008diff
changeset | 341 |   Args.$$$ simpN -- Args.del -- Args.colon >> K (Method.modifier simp_del @{here}),
 | 
| 
aa6296d09e0e
more explicit Method.modifier with reported position;
 wenzelm parents: 
58008diff
changeset | 342 | Args.$$$ simpN -- Args.$$$ onlyN -- Args.colon >> | 
| 
aa6296d09e0e
more explicit Method.modifier with reported position;
 wenzelm parents: 
58008diff
changeset | 343 |     K {init = Raw_Simplifier.clear_simpset, attribute = simp_add, pos = @{here}}]
 | 
| 16014 | 344 | @ cong_modifiers; | 
| 345 | ||
| 346 | val simp_modifiers' = | |
| 58048 
aa6296d09e0e
more explicit Method.modifier with reported position;
 wenzelm parents: 
58008diff
changeset | 347 |  [Args.add -- Args.colon >> K (Method.modifier simp_add @{here}),
 | 
| 
aa6296d09e0e
more explicit Method.modifier with reported position;
 wenzelm parents: 
58008diff
changeset | 348 |   Args.del -- Args.colon >> K (Method.modifier simp_del @{here}),
 | 
| 
aa6296d09e0e
more explicit Method.modifier with reported position;
 wenzelm parents: 
58008diff
changeset | 349 | Args.$$$ onlyN -- Args.colon >> | 
| 
aa6296d09e0e
more explicit Method.modifier with reported position;
 wenzelm parents: 
58008diff
changeset | 350 |     K {init = Raw_Simplifier.clear_simpset, attribute = simp_add, pos = @{here}}]
 | 
| 16014 | 351 | @ cong_modifiers; | 
| 352 | ||
| 31300 | 353 | val simp_options = | 
| 354 | (Args.parens (Args.$$$ no_asmN) >> K simp_tac || | |
| 355 | Args.parens (Args.$$$ no_asm_simpN) >> K asm_simp_tac || | |
| 356 | Args.parens (Args.$$$ no_asm_useN) >> K full_simp_tac || | |
| 357 | Args.parens (Args.$$$ asm_lrN) >> K asm_lr_simp_tac || | |
| 358 | Scan.succeed asm_full_simp_tac); | |
| 16014 | 359 | |
| 31300 | 360 | fun simp_method more_mods meth = | 
| 35613 | 361 | Scan.lift simp_options --| | 
| 31300 | 362 | Method.sections (more_mods @ simp_modifiers') >> | 
| 35613 | 363 | (fn tac => fn ctxt => METHOD (fn facts => meth ctxt tac facts)); | 
| 16014 | 364 | |
| 365 | ||
| 366 | ||
| 18708 | 367 | (** setup **) | 
| 368 | ||
| 31300 | 369 | fun method_setup more_mods = | 
| 56204 | 370 |   Method.setup @{binding simp}
 | 
| 31300 | 371 | (simp_method more_mods (fn ctxt => fn tac => fn facts => | 
| 372 | HEADGOAL (Method.insert_tac facts THEN' | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 373 | (CHANGED_PROP oo tac) ctxt))) | 
| 31300 | 374 | "simplification" #> | 
| 56204 | 375 |   Method.setup @{binding simp_all}
 | 
| 31300 | 376 | (simp_method more_mods (fn ctxt => fn tac => fn facts => | 
| 377 | ALLGOALS (Method.insert_tac facts) THEN | |
| 58008 | 378 | (CHANGED_PROP o PARALLEL_ALLGOALS o tac) ctxt)) | 
| 31300 | 379 | "simplification (all goals)"; | 
| 16014 | 380 | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 381 | fun easy_setup reflect trivs = method_setup [] #> Context.theory_map (map_ss (fn ctxt0 => | 
| 16014 | 382 | let | 
| 383 | val trivialities = Drule.reflexive_thm :: trivs; | |
| 384 | ||
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 385 | fun unsafe_solver_tac ctxt = | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
58963diff
changeset | 386 | FIRST' [resolve_tac ctxt (trivialities @ Raw_Simplifier.prems_of ctxt), assume_tac ctxt]; | 
| 16014 | 387 | val unsafe_solver = mk_solver "easy unsafe" unsafe_solver_tac; | 
| 388 | ||
| 389 | (*no premature instantiation of variables during simplification*) | |
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 390 | fun safe_solver_tac ctxt = | 
| 58957 | 391 | FIRST' [match_tac ctxt (trivialities @ Raw_Simplifier.prems_of ctxt), eq_assume_tac]; | 
| 16014 | 392 | val safe_solver = mk_solver "easy safe" safe_solver_tac; | 
| 393 | ||
| 394 | fun mk_eq thm = | |
| 20872 | 395 | if can Logic.dest_equals (Thm.concl_of thm) then [thm] | 
| 16014 | 396 | else [thm RS reflect] handle THM _ => []; | 
| 397 | ||
| 44058 | 398 | fun mksimps thm = mk_eq (Thm.forall_elim_vars (Thm.maxidx_of thm + 1) thm); | 
| 26497 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 399 | in | 
| 51717 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 wenzelm parents: 
51688diff
changeset | 400 | empty_simpset ctxt0 | 
| 26497 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 401 | setSSolver safe_solver | 
| 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 402 | setSolver unsafe_solver | 
| 45625 
750c5a47400b
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45620diff
changeset | 403 | |> set_subgoaler asm_simp_tac | 
| 
750c5a47400b
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 wenzelm parents: 
45620diff
changeset | 404 | |> set_mksimps (K mksimps) | 
| 26497 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 wenzelm parents: 
26463diff
changeset | 405 | end)); | 
| 16014 | 406 | |
| 407 | end; | |
| 408 | ||
| 32738 | 409 | structure Basic_Simplifier: BASIC_SIMPLIFIER = Simplifier; | 
| 410 | open Basic_Simplifier; |