| author | wenzelm | 
| Tue, 30 Mar 2010 00:47:52 +0200 | |
| changeset 36016 | 4f5c7a19ebe0 | 
| parent 35613 | 9d3ff36ad4e1 | 
| child 36543 | 0e7fc5bf38de | 
| 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  | 
|
5  | 
meta_simplifier.ML for the actual meta-level rewriting engine).  | 
|
6  | 
*)  | 
|
7  | 
||
8  | 
signature BASIC_SIMPLIFIER =  | 
|
9  | 
sig  | 
|
10  | 
include BASIC_META_SIMPLIFIER  | 
|
| 
17883
 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 
wenzelm 
parents: 
17723 
diff
changeset
 | 
11  | 
val change_simpset: (simpset -> simpset) -> unit  | 
| 
32148
 
253f6808dabe
renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
 
wenzelm 
parents: 
32091 
diff
changeset
 | 
12  | 
val global_simpset_of: theory -> simpset  | 
| 16014 | 13  | 
val Addsimprocs: simproc list -> unit  | 
14  | 
val Delsimprocs: simproc list -> unit  | 
|
| 
32148
 
253f6808dabe
renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
 
wenzelm 
parents: 
32091 
diff
changeset
 | 
15  | 
val simpset_of: Proof.context -> simpset  | 
| 
17967
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
16  | 
val generic_simp_tac: bool -> bool * bool * bool -> simpset -> int -> tactic  | 
| 16014 | 17  | 
val safe_asm_full_simp_tac: simpset -> int -> tactic  | 
| 35613 | 18  | 
val simp_tac: simpset -> int -> tactic  | 
19  | 
val asm_simp_tac: simpset -> int -> tactic  | 
|
20  | 
val full_simp_tac: simpset -> int -> tactic  | 
|
21  | 
val asm_lr_simp_tac: simpset -> int -> tactic  | 
|
22  | 
val asm_full_simp_tac: simpset -> int -> tactic  | 
|
23  | 
val simplify: simpset -> thm -> thm  | 
|
24  | 
val asm_simplify: simpset -> thm -> thm  | 
|
25  | 
val full_simplify: simpset -> thm -> thm  | 
|
26  | 
val asm_lr_simplify: simpset -> thm -> thm  | 
|
| 16014 | 27  | 
val asm_full_simplify: simpset -> thm -> thm  | 
28  | 
end;  | 
|
29  | 
||
30  | 
signature SIMPLIFIER =  | 
|
31  | 
sig  | 
|
32  | 
include BASIC_SIMPLIFIER  | 
|
| 30356 | 33  | 
val pretty_ss: Proof.context -> simpset -> Pretty.T  | 
| 17004 | 34  | 
val clear_ss: simpset -> simpset  | 
| 32738 | 35  | 
val debug_bounds: bool Unsynchronized.ref  | 
| 
17883
 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 
wenzelm 
parents: 
17723 
diff
changeset
 | 
36  | 
val inherit_context: simpset -> simpset -> simpset  | 
| 17898 | 37  | 
val the_context: simpset -> Proof.context  | 
38  | 
val context: Proof.context -> simpset -> simpset  | 
|
| 
35232
 
f588e1169c8b
renamed Simplifier.theory_context to Simplifier.global_context to emphasize that this is not the real thing;
 
wenzelm 
parents: 
33671 
diff
changeset
 | 
39  | 
val global_context: theory -> simpset -> simpset  | 
| 16458 | 40  | 
val simproc_i: theory -> string -> term list  | 
41  | 
-> (theory -> simpset -> term -> thm option) -> simproc  | 
|
42  | 
val simproc: theory -> string -> string list  | 
|
43  | 
-> (theory -> simpset -> term -> thm option) -> simproc  | 
|
| 35613 | 44  | 
val rewrite: simpset -> conv  | 
45  | 
val asm_rewrite: simpset -> conv  | 
|
46  | 
val full_rewrite: simpset -> conv  | 
|
47  | 
val asm_lr_rewrite: simpset -> conv  | 
|
| 23598 | 48  | 
val asm_full_rewrite: simpset -> conv  | 
| 22379 | 49  | 
val get_ss: Context.generic -> simpset  | 
50  | 
val map_ss: (simpset -> simpset) -> Context.generic -> Context.generic  | 
|
| 18728 | 51  | 
val attrib: (simpset * thm list -> simpset) -> attribute  | 
52  | 
val simp_add: attribute  | 
|
53  | 
val simp_del: attribute  | 
|
54  | 
val cong_add: attribute  | 
|
55  | 
val cong_del: attribute  | 
|
| 
26497
 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
56  | 
val map_simpset: (simpset -> simpset) -> theory -> theory  | 
| 24024 | 57  | 
val get_simproc: Context.generic -> xstring -> simproc  | 
| 22236 | 58  | 
  val def_simproc: {name: string, lhss: string list,
 | 
59  | 
proc: morphism -> simpset -> cterm -> thm option, identifier: thm list} ->  | 
|
| 22201 | 60  | 
local_theory -> local_theory  | 
| 22236 | 61  | 
  val def_simproc_i: {name: string, lhss: term list,
 | 
62  | 
proc: morphism -> simpset -> cterm -> thm option, identifier: thm list} ->  | 
|
| 22201 | 63  | 
local_theory -> local_theory  | 
| 30513 | 64  | 
val cong_modifiers: Method.modifier parser list  | 
65  | 
val simp_modifiers': Method.modifier parser list  | 
|
66  | 
val simp_modifiers: Method.modifier parser list  | 
|
67  | 
val method_setup: Method.modifier parser list -> theory -> theory  | 
|
| 18708 | 68  | 
val easy_setup: thm -> thm list -> theory -> theory  | 
| 16014 | 69  | 
end;  | 
70  | 
||
71  | 
structure Simplifier: SIMPLIFIER =  | 
|
72  | 
struct  | 
|
73  | 
||
| 21708 | 74  | 
open MetaSimplifier;  | 
75  | 
||
76  | 
||
| 30356 | 77  | 
(** pretty printing **)  | 
78  | 
||
79  | 
fun pretty_ss ctxt ss =  | 
|
80  | 
let  | 
|
81  | 
val pretty_cterm = Syntax.pretty_term ctxt o Thm.term_of;  | 
|
| 
32091
 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 
wenzelm 
parents: 
31300 
diff
changeset
 | 
82  | 
val pretty_thm = Display.pretty_thm ctxt;  | 
| 30356 | 83  | 
fun pretty_proc (name, lhss) = Pretty.big_list (name ^ ":") (map pretty_cterm lhss);  | 
84  | 
fun pretty_cong (name, thm) =  | 
|
85  | 
Pretty.block [Pretty.str (name ^ ":"), Pretty.brk 1, pretty_thm thm];  | 
|
86  | 
||
87  | 
    val {simps, procs, congs, loopers, unsafe_solvers, safe_solvers, ...} = dest_ss ss;
 | 
|
88  | 
in  | 
|
89  | 
[Pretty.big_list "simplification rules:" (map (pretty_thm o #2) simps),  | 
|
90  | 
Pretty.big_list "simplification procedures:" (map pretty_proc (sort_wrt #1 procs)),  | 
|
91  | 
Pretty.big_list "congruences:" (map pretty_cong congs),  | 
|
92  | 
      Pretty.strs ("loopers:" :: map quote loopers),
 | 
|
93  | 
      Pretty.strs ("unsafe solvers:" :: map quote unsafe_solvers),
 | 
|
94  | 
      Pretty.strs ("safe solvers:" :: map quote safe_solvers)]
 | 
|
95  | 
|> Pretty.chunks  | 
|
96  | 
end;  | 
|
97  | 
||
98  | 
||
99  | 
||
| 
17883
 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 
wenzelm 
parents: 
17723 
diff
changeset
 | 
100  | 
(** simpset data **)  | 
| 16014 | 101  | 
|
| 33519 | 102  | 
structure SimpsetData = Generic_Data  | 
| 22846 | 103  | 
(  | 
| 
26497
 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
104  | 
type T = simpset;  | 
| 
 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
105  | 
val empty = empty_ss;  | 
| 
 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
106  | 
fun extend ss = MetaSimplifier.inherit_context empty_ss ss;  | 
| 33519 | 107  | 
val merge = merge_ss;  | 
| 22846 | 108  | 
);  | 
| 16014 | 109  | 
|
| 
26497
 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
110  | 
val get_ss = SimpsetData.get;  | 
| 
 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
111  | 
val map_ss = SimpsetData.map;  | 
| 
 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
112  | 
|
| 
 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
113  | 
|
| 
 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
114  | 
(* attributes *)  | 
| 
 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
115  | 
|
| 
 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
116  | 
fun attrib f = Thm.declaration_attribute (fn th => map_ss (fn ss => f (ss, [th])));  | 
| 16014 | 117  | 
|
| 
26497
 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
118  | 
val simp_add = attrib (op addsimps);  | 
| 
 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
119  | 
val simp_del = attrib (op delsimps);  | 
| 
 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
120  | 
val cong_add = attrib (op addcongs);  | 
| 
 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
121  | 
val cong_del = attrib (op delcongs);  | 
| 
 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
122  | 
|
| 16014 | 123  | 
|
| 
26497
 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
124  | 
(* global simpset *)  | 
| 
 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
125  | 
|
| 
 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
126  | 
fun map_simpset f = Context.theory_map (map_ss f);  | 
| 
 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
127  | 
fun change_simpset f = Context.>> (Context.map_theory (map_simpset f));  | 
| 
32148
 
253f6808dabe
renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
 
wenzelm 
parents: 
32091 
diff
changeset
 | 
128  | 
fun global_simpset_of thy =  | 
| 
 
253f6808dabe
renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
 
wenzelm 
parents: 
32091 
diff
changeset
 | 
129  | 
MetaSimplifier.context (ProofContext.init thy) (get_ss (Context.Theory thy));  | 
| 16014 | 130  | 
|
| 
17883
 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 
wenzelm 
parents: 
17723 
diff
changeset
 | 
131  | 
fun Addsimprocs args = change_simpset (fn ss => ss addsimprocs args);  | 
| 
 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 
wenzelm 
parents: 
17723 
diff
changeset
 | 
132  | 
fun Delsimprocs args = change_simpset (fn ss => ss delsimprocs args);  | 
| 16014 | 133  | 
|
134  | 
||
| 
26497
 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
135  | 
(* local simpset *)  | 
| 16014 | 136  | 
|
| 
32148
 
253f6808dabe
renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
 
wenzelm 
parents: 
32091 
diff
changeset
 | 
137  | 
fun simpset_of ctxt = MetaSimplifier.context ctxt (get_ss (Context.Proof ctxt));  | 
| 16014 | 138  | 
|
| 27338 | 139  | 
val _ = ML_Antiquote.value "simpset"  | 
| 
32148
 
253f6808dabe
renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
 
wenzelm 
parents: 
32091 
diff
changeset
 | 
140  | 
(Scan.succeed "Simplifier.simpset_of (ML_Context.the_local_context ())");  | 
| 22132 | 141  | 
|
| 16014 | 142  | 
|
| 
17967
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
143  | 
|
| 22201 | 144  | 
(** named simprocs **)  | 
145  | 
||
| 22204 | 146  | 
(* data *)  | 
147  | 
||
| 33519 | 148  | 
structure Simprocs = Generic_Data  | 
| 22201 | 149  | 
(  | 
| 
33095
 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 
wenzelm 
parents: 
33092 
diff
changeset
 | 
150  | 
type T = simproc Name_Space.table;  | 
| 33159 | 151  | 
val empty : T = Name_Space.empty_table "simproc";  | 
| 22201 | 152  | 
val extend = I;  | 
| 33519 | 153  | 
fun merge simprocs = Name_Space.merge_tables simprocs;  | 
| 22201 | 154  | 
);  | 
155  | 
||
| 22204 | 156  | 
|
157  | 
(* get simprocs *)  | 
|
158  | 
||
| 24024 | 159  | 
fun get_simproc context xname =  | 
| 22204 | 160  | 
let  | 
| 24024 | 161  | 
val (space, tab) = Simprocs.get context;  | 
| 
33095
 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 
wenzelm 
parents: 
33092 
diff
changeset
 | 
162  | 
val name = Name_Space.intern space xname;  | 
| 22204 | 163  | 
in  | 
164  | 
(case Symtab.lookup tab name of  | 
|
| 22236 | 165  | 
SOME proc => proc  | 
| 22204 | 166  | 
    | NONE => error ("Undefined simplification procedure: " ^ quote name))
 | 
167  | 
end;  | 
|
168  | 
||
| 27338 | 169  | 
val _ = ML_Antiquote.value "simproc" (Scan.lift Args.name >> (fn name =>  | 
170  | 
"Simplifier.get_simproc (ML_Context.the_generic_context ()) " ^ ML_Syntax.print_string name));  | 
|
| 22204 | 171  | 
|
172  | 
||
173  | 
(* define simprocs *)  | 
|
| 22201 | 174  | 
|
175  | 
local  | 
|
176  | 
||
| 22236 | 177  | 
fun gen_simproc prep {name, lhss, proc, identifier} lthy =
 | 
| 22201 | 178  | 
let  | 
| 28991 | 179  | 
val b = Binding.name name;  | 
| 33671 | 180  | 
val naming = Local_Theory.naming_of lthy;  | 
| 22236 | 181  | 
val simproc = make_simproc  | 
| 33169 | 182  | 
      {name = Name_Space.full_name naming b,
 | 
| 22236 | 183  | 
lhss =  | 
184  | 
let  | 
|
185  | 
val lhss' = prep lthy lhss;  | 
|
186  | 
val ctxt' = lthy  | 
|
187  | 
|> fold Variable.declare_term lhss'  | 
|
188  | 
|> fold Variable.auto_fixes lhss';  | 
|
189  | 
in Variable.export_terms ctxt' lthy lhss' end  | 
|
190  | 
|> map (Thm.cterm_of (ProofContext.theory_of lthy)),  | 
|
191  | 
proc = proc,  | 
|
| 
33551
 
c40ced05b10a
define simprocs: do not apply target_morphism prematurely, this is already done in LocalTheory.declaration;
 
wenzelm 
parents: 
33519 
diff
changeset
 | 
192  | 
identifier = identifier};  | 
| 22201 | 193  | 
in  | 
| 33671 | 194  | 
lthy |> Local_Theory.declaration false (fn phi =>  | 
| 22201 | 195  | 
let  | 
| 28991 | 196  | 
val b' = Morphism.binding phi b;  | 
| 22236 | 197  | 
val simproc' = morph_simproc phi simproc;  | 
| 22201 | 198  | 
in  | 
| 
33095
 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 
wenzelm 
parents: 
33092 
diff
changeset
 | 
199  | 
Simprocs.map (#2 o Name_Space.define true naming (b', simproc'))  | 
| 24024 | 200  | 
#> map_ss (fn ss => ss addsimprocs [simproc'])  | 
| 22201 | 201  | 
end)  | 
202  | 
end;  | 
|
203  | 
||
204  | 
in  | 
|
205  | 
||
| 
24509
 
23ee6b7788c2
replaced ProofContext.read_term/prop by general Syntax.read_term/prop;
 
wenzelm 
parents: 
24124 
diff
changeset
 | 
206  | 
val def_simproc = gen_simproc Syntax.read_terms;  | 
| 
 
23ee6b7788c2
replaced ProofContext.read_term/prop by general Syntax.read_term/prop;
 
wenzelm 
parents: 
24124 
diff
changeset
 | 
207  | 
val def_simproc_i = gen_simproc Syntax.check_terms;  | 
| 22201 | 208  | 
|
209  | 
end;  | 
|
210  | 
||
211  | 
||
212  | 
||
| 
17967
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
213  | 
(** simplification tactics and rules **)  | 
| 
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
214  | 
|
| 
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
215  | 
fun solve_all_tac solvers ss =  | 
| 
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
216  | 
let  | 
| 30336 | 217  | 
    val (_, {subgoal_tac, ...}) = MetaSimplifier.internal_ss ss;
 | 
| 
17967
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
218  | 
val solve_tac = subgoal_tac (MetaSimplifier.set_solvers solvers ss) THEN_ALL_NEW (K no_tac);  | 
| 
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
219  | 
in DEPTH_SOLVE (solve_tac 1) end;  | 
| 
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
220  | 
|
| 
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
221  | 
(*NOTE: may instantiate unknowns that appear also in other subgoals*)  | 
| 
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
222  | 
fun generic_simp_tac safe mode ss =  | 
| 
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
223  | 
let  | 
| 30336 | 224  | 
    val (_, {loop_tacs, solvers = (unsafe_solvers, solvers), ...}) = MetaSimplifier.internal_ss ss;
 | 
| 
21286
 
b5e7b80caa6a
introduces canonical AList functions for loop_tacs
 
haftmann 
parents: 
20872 
diff
changeset
 | 
225  | 
val loop_tac = FIRST' (map (fn (_, tac) => tac ss) (rev loop_tacs));  | 
| 
17967
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
226  | 
val solve_tac = FIRST' (map (MetaSimplifier.solver ss)  | 
| 22717 | 227  | 
(rev (if safe then solvers else unsafe_solvers)));  | 
| 
17967
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
228  | 
|
| 
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
229  | 
fun simp_loop_tac i =  | 
| 
23536
 
60a1672e298e
moved (asm_)rewrite_goal_tac from goal.ML to meta_simplifier.ML (no longer depends on SELECT_GOAL);
 
wenzelm 
parents: 
23086 
diff
changeset
 | 
230  | 
asm_rewrite_goal_tac mode (solve_all_tac unsafe_solvers) ss i THEN  | 
| 
17967
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
231  | 
(solve_tac i ORELSE TRY ((loop_tac THEN_ALL_NEW simp_loop_tac) i));  | 
| 
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
232  | 
in simp_loop_tac end;  | 
| 
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
233  | 
|
| 
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
234  | 
local  | 
| 
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
235  | 
|
| 
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
236  | 
fun simp rew mode ss thm =  | 
| 
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
237  | 
let  | 
| 30336 | 238  | 
    val (_, {solvers = (unsafe_solvers, _), ...}) = MetaSimplifier.internal_ss ss;
 | 
| 22717 | 239  | 
val tacf = solve_all_tac (rev unsafe_solvers);  | 
| 
17967
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
240  | 
fun prover s th = Option.map #1 (Seq.pull (tacf s th));  | 
| 
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
241  | 
in rew mode prover ss thm end;  | 
| 
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
242  | 
|
| 
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
243  | 
in  | 
| 
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
244  | 
|
| 
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
245  | 
val simp_thm = simp MetaSimplifier.rewrite_thm;  | 
| 
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
246  | 
val simp_cterm = simp MetaSimplifier.rewrite_cterm;  | 
| 
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
247  | 
|
| 
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
248  | 
end;  | 
| 
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
249  | 
|
| 
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
250  | 
|
| 16806 | 251  | 
(* tactics *)  | 
252  | 
||
| 16014 | 253  | 
val simp_tac = generic_simp_tac false (false, false, false);  | 
254  | 
val asm_simp_tac = generic_simp_tac false (false, true, false);  | 
|
255  | 
val full_simp_tac = generic_simp_tac false (true, false, false);  | 
|
256  | 
val asm_lr_simp_tac = generic_simp_tac false (true, true, false);  | 
|
257  | 
val asm_full_simp_tac = generic_simp_tac false (true, true, true);  | 
|
258  | 
val safe_asm_full_simp_tac = generic_simp_tac true (true, true, true);  | 
|
259  | 
||
| 16806 | 260  | 
|
261  | 
(* conversions *)  | 
|
262  | 
||
| 
17967
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
263  | 
val simplify = simp_thm (false, false, false);  | 
| 
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
264  | 
val asm_simplify = simp_thm (false, true, false);  | 
| 
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
265  | 
val full_simplify = simp_thm (true, false, false);  | 
| 
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
266  | 
val asm_lr_simplify = simp_thm (true, true, false);  | 
| 
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
267  | 
val asm_full_simplify = simp_thm (true, true, true);  | 
| 16014 | 268  | 
|
| 
17967
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
269  | 
val rewrite = simp_cterm (false, false, false);  | 
| 
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
270  | 
val asm_rewrite = simp_cterm (false, true, false);  | 
| 
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
271  | 
val full_rewrite = simp_cterm (true, false, false);  | 
| 
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
272  | 
val asm_lr_rewrite = simp_cterm (true, true, false);  | 
| 
 
7a733b7438e1
added simplification tactics and rules (from meta_simplifier.ML);
 
wenzelm 
parents: 
17898 
diff
changeset
 | 
273  | 
val asm_full_rewrite = simp_cterm (true, true, true);  | 
| 16014 | 274  | 
|
275  | 
||
276  | 
||
277  | 
(** concrete syntax of attributes **)  | 
|
278  | 
||
279  | 
(* add / del *)  | 
|
280  | 
||
281  | 
val simpN = "simp";  | 
|
282  | 
val congN = "cong";  | 
|
283  | 
val onlyN = "only";  | 
|
284  | 
val no_asmN = "no_asm";  | 
|
285  | 
val no_asm_useN = "no_asm_use";  | 
|
286  | 
val no_asm_simpN = "no_asm_simp";  | 
|
287  | 
val asm_lrN = "asm_lr";  | 
|
288  | 
||
289  | 
||
| 24024 | 290  | 
(* simprocs *)  | 
291  | 
||
292  | 
local  | 
|
293  | 
||
294  | 
val add_del =  | 
|
295  | 
(Args.del -- Args.colon >> K (op delsimprocs) ||  | 
|
296  | 
Scan.option (Args.add -- Args.colon) >> K (op addsimprocs))  | 
|
297  | 
>> (fn f => fn simproc => fn phi => Thm.declaration_attribute  | 
|
298  | 
(K (map_ss (fn ss => f (ss, [morph_simproc phi simproc])))));  | 
|
299  | 
||
300  | 
in  | 
|
301  | 
||
| 30528 | 302  | 
val simproc_att =  | 
303  | 
Scan.peek (fn context =>  | 
|
| 24024 | 304  | 
add_del :|-- (fn decl =>  | 
305  | 
Scan.repeat1 (Args.named_attribute (decl o get_simproc context))  | 
|
| 30528 | 306  | 
>> (Library.apply o map Morphism.form)));  | 
| 24024 | 307  | 
|
308  | 
end;  | 
|
| 
24124
 
4399175e3014
turned simp_depth_limit into configuration option;
 
wenzelm 
parents: 
24024 
diff
changeset
 | 
309  | 
|
| 24024 | 310  | 
|
| 16014 | 311  | 
(* conversions *)  | 
312  | 
||
313  | 
local  | 
|
314  | 
||
315  | 
fun conv_mode x =  | 
|
316  | 
((Args.parens (Args.$$$ no_asmN) >> K simplify ||  | 
|
317  | 
Args.parens (Args.$$$ no_asm_simpN) >> K asm_simplify ||  | 
|
318  | 
Args.parens (Args.$$$ no_asm_useN) >> K full_simplify ||  | 
|
319  | 
Scan.succeed asm_full_simplify) |> Scan.lift) x;  | 
|
320  | 
||
321  | 
in  | 
|
322  | 
||
| 30528 | 323  | 
val simplified = conv_mode -- Attrib.thms >>  | 
324  | 
(fn (f, ths) => Thm.rule_attribute (fn context =>  | 
|
| 27021 | 325  | 
f ((if null ths then I else MetaSimplifier.clear_ss)  | 
| 
32148
 
253f6808dabe
renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
 
wenzelm 
parents: 
32091 
diff
changeset
 | 
326  | 
(simpset_of (Context.proof_of context)) addsimps ths)));  | 
| 16014 | 327  | 
|
328  | 
end;  | 
|
329  | 
||
330  | 
||
331  | 
(* setup attributes *)  | 
|
332  | 
||
| 26463 | 333  | 
val _ = Context.>> (Context.map_theory  | 
| 30528 | 334  | 
(Attrib.setup (Binding.name simpN) (Attrib.add_del simp_add simp_del)  | 
335  | 
"declaration of Simplifier rewrite rule" #>  | 
|
336  | 
Attrib.setup (Binding.name congN) (Attrib.add_del cong_add cong_del)  | 
|
337  | 
"declaration of Simplifier congruence rule" #>  | 
|
| 33671 | 338  | 
Attrib.setup (Binding.name "simproc") simproc_att  | 
339  | 
"declaration of simplification procedures" #>  | 
|
| 30528 | 340  | 
Attrib.setup (Binding.name "simplified") simplified "simplified rule"));  | 
| 16014 | 341  | 
|
342  | 
||
343  | 
||
| 31300 | 344  | 
(** method syntax **)  | 
| 16014 | 345  | 
|
346  | 
val cong_modifiers =  | 
|
| 18728 | 347  | 
[Args.$$$ congN -- Args.colon >> K ((I, cong_add): Method.modifier),  | 
348  | 
Args.$$$ congN -- Args.add -- Args.colon >> K (I, cong_add),  | 
|
349  | 
Args.$$$ congN -- Args.del -- Args.colon >> K (I, cong_del)];  | 
|
| 16014 | 350  | 
|
351  | 
val simp_modifiers =  | 
|
| 18728 | 352  | 
[Args.$$$ simpN -- Args.colon >> K (I, simp_add),  | 
353  | 
Args.$$$ simpN -- Args.add -- Args.colon >> K (I, simp_add),  | 
|
354  | 
Args.$$$ simpN -- Args.del -- Args.colon >> K (I, simp_del),  | 
|
| 
17883
 
efa1bc2bdcc6
removed obsolete/experimental context components (superceded by Simplifier.the_context);
 
wenzelm 
parents: 
17723 
diff
changeset
 | 
355  | 
Args.$$$ simpN -- Args.$$$ onlyN -- Args.colon  | 
| 
26497
 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
356  | 
>> K (Context.proof_map (map_ss MetaSimplifier.clear_ss), simp_add)]  | 
| 16014 | 357  | 
@ cong_modifiers;  | 
358  | 
||
359  | 
val simp_modifiers' =  | 
|
| 18728 | 360  | 
[Args.add -- Args.colon >> K (I, simp_add),  | 
361  | 
Args.del -- Args.colon >> K (I, simp_del),  | 
|
| 18688 | 362  | 
Args.$$$ onlyN -- Args.colon  | 
| 
26497
 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
363  | 
>> K (Context.proof_map (map_ss MetaSimplifier.clear_ss), simp_add)]  | 
| 16014 | 364  | 
@ cong_modifiers;  | 
365  | 
||
| 31300 | 366  | 
val simp_options =  | 
367  | 
(Args.parens (Args.$$$ no_asmN) >> K simp_tac ||  | 
|
368  | 
Args.parens (Args.$$$ no_asm_simpN) >> K asm_simp_tac ||  | 
|
369  | 
Args.parens (Args.$$$ no_asm_useN) >> K full_simp_tac ||  | 
|
370  | 
Args.parens (Args.$$$ asm_lrN) >> K asm_lr_simp_tac ||  | 
|
371  | 
Scan.succeed asm_full_simp_tac);  | 
|
| 16014 | 372  | 
|
| 31300 | 373  | 
fun simp_method more_mods meth =  | 
| 35613 | 374  | 
Scan.lift simp_options --|  | 
| 31300 | 375  | 
Method.sections (more_mods @ simp_modifiers') >>  | 
| 35613 | 376  | 
(fn tac => fn ctxt => METHOD (fn facts => meth ctxt tac facts));  | 
| 16014 | 377  | 
|
378  | 
||
379  | 
||
| 18708 | 380  | 
(** setup **)  | 
381  | 
||
| 31300 | 382  | 
fun method_setup more_mods =  | 
383  | 
Method.setup (Binding.name simpN)  | 
|
384  | 
(simp_method more_mods (fn ctxt => fn tac => fn facts =>  | 
|
385  | 
HEADGOAL (Method.insert_tac facts THEN'  | 
|
| 
32148
 
253f6808dabe
renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
 
wenzelm 
parents: 
32091 
diff
changeset
 | 
386  | 
(CHANGED_PROP oo tac) (simpset_of ctxt))))  | 
| 31300 | 387  | 
"simplification" #>  | 
388  | 
Method.setup (Binding.name "simp_all")  | 
|
389  | 
(simp_method more_mods (fn ctxt => fn tac => fn facts =>  | 
|
390  | 
ALLGOALS (Method.insert_tac facts) THEN  | 
|
| 
32148
 
253f6808dabe
renamed simpset_of to global_simpset_of, and local_simpset_of to simpset_of -- same for claset and clasimpset;
 
wenzelm 
parents: 
32091 
diff
changeset
 | 
391  | 
(CHANGED_PROP o ALLGOALS o tac) (simpset_of ctxt)))  | 
| 31300 | 392  | 
"simplification (all goals)";  | 
| 16014 | 393  | 
|
| 
26497
 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
394  | 
fun easy_setup reflect trivs = method_setup [] #> Context.theory_map (map_ss (fn _ =>  | 
| 16014 | 395  | 
let  | 
396  | 
val trivialities = Drule.reflexive_thm :: trivs;  | 
|
397  | 
||
398  | 
fun unsafe_solver_tac prems = FIRST' [resolve_tac (trivialities @ prems), assume_tac];  | 
|
399  | 
val unsafe_solver = mk_solver "easy unsafe" unsafe_solver_tac;  | 
|
400  | 
||
401  | 
(*no premature instantiation of variables during simplification*)  | 
|
402  | 
fun safe_solver_tac prems = FIRST' [match_tac (trivialities @ prems), eq_assume_tac];  | 
|
403  | 
val safe_solver = mk_solver "easy safe" safe_solver_tac;  | 
|
404  | 
||
405  | 
fun mk_eq thm =  | 
|
| 20872 | 406  | 
if can Logic.dest_equals (Thm.concl_of thm) then [thm]  | 
| 16014 | 407  | 
else [thm RS reflect] handle THM _ => [];  | 
408  | 
||
| 26653 | 409  | 
fun mksimps thm = mk_eq (Thm.forall_elim_vars (#maxidx (Thm.rep_thm thm) + 1) thm);  | 
| 
26497
 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
410  | 
in  | 
| 
 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
411  | 
empty_ss setsubgoaler asm_simp_tac  | 
| 
 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
412  | 
setSSolver safe_solver  | 
| 
 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
413  | 
setSolver unsafe_solver  | 
| 
 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
414  | 
setmksimps mksimps  | 
| 
 
1873915c64a9
purely functional setup of claset/simpset/clasimpset;
 
wenzelm 
parents: 
26463 
diff
changeset
 | 
415  | 
end));  | 
| 16014 | 416  | 
|
417  | 
end;  | 
|
418  | 
||
| 32738 | 419  | 
structure Basic_Simplifier: BASIC_SIMPLIFIER = Simplifier;  | 
420  | 
open Basic_Simplifier;  |