| author | blanchet | 
| Wed, 28 May 2014 09:44:14 +0200 | |
| changeset 57098 | c0a25c7c4b8e | 
| parent 56438 | 7f6b2634d853 | 
| child 57859 | 29e728588163 | 
| permissions | -rw-r--r-- | 
| 
41228
 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 
wenzelm 
parents: 
41227 
diff
changeset
 | 
1  | 
(* Title: Pure/raw_simplifier.ML  | 
| 
29269
 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 
wenzelm 
parents: 
28839 
diff
changeset
 | 
2  | 
Author: Tobias Nipkow and Stefan Berghofer, TU Muenchen  | 
| 10413 | 3  | 
|
| 
41228
 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 
wenzelm 
parents: 
41227 
diff
changeset
 | 
4  | 
Higher-order Simplification.  | 
| 10413 | 5  | 
*)  | 
6  | 
||
| 
15006
 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 
skalberg 
parents: 
15001 
diff
changeset
 | 
7  | 
infix 4  | 
| 
45620
 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 
wenzelm 
parents: 
45405 
diff
changeset
 | 
8  | 
addsimps delsimps addsimprocs delsimprocs  | 
| 52037 | 9  | 
setloop addloop delloop  | 
| 
45625
 
750c5a47400b
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 
wenzelm 
parents: 
45621 
diff
changeset
 | 
10  | 
setSSolver addSSolver setSolver addSolver;  | 
| 
15006
 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 
skalberg 
parents: 
15001 
diff
changeset
 | 
11  | 
|
| 
41228
 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 
wenzelm 
parents: 
41227 
diff
changeset
 | 
12  | 
signature BASIC_RAW_SIMPLIFIER =  | 
| 11672 | 13  | 
sig  | 
| 41227 | 14  | 
val simp_depth_limit: int Config.T  | 
15  | 
val simp_trace_depth_limit: int Config.T  | 
|
| 
40878
 
7695e4de4d86
renamed trace_simp to simp_trace, and debug_simp to simp_debug;
 
wenzelm 
parents: 
39163 
diff
changeset
 | 
16  | 
val simp_debug: bool Config.T  | 
| 
 
7695e4de4d86
renamed trace_simp to simp_trace, and debug_simp to simp_debug;
 
wenzelm 
parents: 
39163 
diff
changeset
 | 
17  | 
val simp_trace: bool Config.T  | 
| 51590 | 18  | 
type cong_name = bool * string  | 
| 15023 | 19  | 
type rrule  | 
| 
55316
 
885500f4aa6a
interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
 
Lars Hupel <lars.hupel@mytum.de> 
parents: 
55032 
diff
changeset
 | 
20  | 
val mk_rrules: Proof.context -> thm list -> rrule list  | 
| 16807 | 21  | 
val eq_rrule: rrule * rrule -> bool  | 
| 15023 | 22  | 
type proc  | 
| 17614 | 23  | 
type solver  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
24  | 
val mk_solver: string -> (Proof.context -> int -> tactic) -> solver  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
25  | 
type simpset  | 
| 15023 | 26  | 
val empty_ss: simpset  | 
27  | 
val merge_ss: simpset * simpset -> simpset  | 
|
| 30356 | 28  | 
val dest_ss: simpset ->  | 
29  | 
   {simps: (string * thm) list,
 | 
|
30  | 
procs: (string * cterm list) list,  | 
|
| 51590 | 31  | 
congs: (cong_name * thm) list,  | 
32  | 
weak_congs: cong_name list,  | 
|
| 30356 | 33  | 
loopers: string list,  | 
34  | 
unsafe_solvers: string list,  | 
|
35  | 
safe_solvers: string list}  | 
|
| 15023 | 36  | 
type simproc  | 
| 22234 | 37  | 
val eq_simproc: simproc * simproc -> bool  | 
| 45290 | 38  | 
val transform_simproc: morphism -> simproc -> simproc  | 
| 22234 | 39  | 
  val make_simproc: {name: string, lhss: cterm list,
 | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
40  | 
proc: morphism -> Proof.context -> cterm -> thm option, identifier: thm list} -> simproc  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
41  | 
val mk_simproc: string -> cterm list -> (Proof.context -> term -> thm option) -> simproc  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
42  | 
val simpset_of: Proof.context -> simpset  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
43  | 
val put_simpset: simpset -> Proof.context -> Proof.context  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
44  | 
val simpset_map: Proof.context -> (Proof.context -> Proof.context) -> simpset -> simpset  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
45  | 
val map_theory_simpset: (Proof.context -> Proof.context) -> theory -> theory  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
46  | 
val empty_simpset: Proof.context -> Proof.context  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
47  | 
val clear_simpset: Proof.context -> Proof.context  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
48  | 
val addsimps: Proof.context * thm list -> Proof.context  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
49  | 
val delsimps: Proof.context * thm list -> Proof.context  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
50  | 
val addsimprocs: Proof.context * simproc list -> Proof.context  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
51  | 
val delsimprocs: Proof.context * simproc list -> Proof.context  | 
| 52037 | 52  | 
val setloop: Proof.context * (Proof.context -> int -> tactic) -> Proof.context  | 
53  | 
val addloop: Proof.context * (string * (Proof.context -> int -> tactic)) -> Proof.context  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
54  | 
val delloop: Proof.context * string -> Proof.context  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
55  | 
val setSSolver: Proof.context * solver -> Proof.context  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
56  | 
val addSSolver: Proof.context * solver -> Proof.context  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
57  | 
val setSolver: Proof.context * solver -> Proof.context  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
58  | 
val addSolver: Proof.context * solver -> Proof.context  | 
| 21708 | 59  | 
|
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
54731 
diff
changeset
 | 
60  | 
val rewrite_rule: Proof.context -> thm list -> thm -> thm  | 
| 
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
54731 
diff
changeset
 | 
61  | 
val rewrite_goals_rule: Proof.context -> thm list -> thm -> thm  | 
| 
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
54731 
diff
changeset
 | 
62  | 
val rewrite_goals_tac: Proof.context -> thm list -> tactic  | 
| 
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
54731 
diff
changeset
 | 
63  | 
val rewrite_goal_tac: Proof.context -> thm list -> int -> tactic  | 
| 
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
54731 
diff
changeset
 | 
64  | 
val prune_params_tac: Proof.context -> tactic  | 
| 
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
54731 
diff
changeset
 | 
65  | 
val fold_rule: Proof.context -> thm list -> thm -> thm  | 
| 
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
54731 
diff
changeset
 | 
66  | 
val fold_goals_tac: Proof.context -> thm list -> tactic  | 
| 
54883
 
dd04a8b654fc
proper context for norm_hhf and derived operations;
 
wenzelm 
parents: 
54742 
diff
changeset
 | 
67  | 
val norm_hhf: Proof.context -> thm -> thm  | 
| 
 
dd04a8b654fc
proper context for norm_hhf and derived operations;
 
wenzelm 
parents: 
54742 
diff
changeset
 | 
68  | 
val norm_hhf_protect: Proof.context -> thm -> thm  | 
| 
15006
 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 
skalberg 
parents: 
15001 
diff
changeset
 | 
69  | 
end;  | 
| 
 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 
skalberg 
parents: 
15001 
diff
changeset
 | 
70  | 
|
| 
41228
 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 
wenzelm 
parents: 
41227 
diff
changeset
 | 
71  | 
signature RAW_SIMPLIFIER =  | 
| 10413 | 72  | 
sig  | 
| 
41228
 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 
wenzelm 
parents: 
41227 
diff
changeset
 | 
73  | 
include BASIC_RAW_SIMPLIFIER  | 
| 54997 | 74  | 
exception SIMPLIFIER of string * thm list  | 
| 
54729
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
75  | 
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: 
54729 
diff
changeset
 | 
76  | 
val set_trace_ops: trace_ops -> theory -> theory  | 
| 30336 | 77  | 
val internal_ss: simpset ->  | 
| 51590 | 78  | 
   {congs: (cong_name * thm) list * cong_name list,
 | 
| 30336 | 79  | 
procs: proc Net.net,  | 
80  | 
mk_rews:  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
81  | 
     {mk: Proof.context -> thm -> thm list,
 | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
82  | 
mk_cong: Proof.context -> thm -> thm,  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
83  | 
mk_sym: Proof.context -> thm -> thm option,  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
84  | 
mk_eq_True: Proof.context -> thm -> thm option,  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
85  | 
reorient: Proof.context -> term list -> term -> term -> bool},  | 
| 30336 | 86  | 
termless: term * term -> bool,  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
87  | 
subgoal_tac: Proof.context -> int -> tactic,  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
88  | 
loop_tacs: (string * (Proof.context -> int -> tactic)) list,  | 
| 
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: 
54729 
diff
changeset
 | 
89  | 
solvers: solver list * solver list}  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
90  | 
val map_ss: (Proof.context -> Proof.context) -> Context.generic -> Context.generic  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
91  | 
val prems_of: Proof.context -> thm list  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
92  | 
val add_simp: thm -> Proof.context -> Proof.context  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
93  | 
val del_simp: thm -> Proof.context -> Proof.context  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
94  | 
val add_eqcong: thm -> Proof.context -> Proof.context  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
95  | 
val del_eqcong: thm -> Proof.context -> Proof.context  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
96  | 
val add_cong: thm -> Proof.context -> Proof.context  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
97  | 
val del_cong: thm -> Proof.context -> Proof.context  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
98  | 
val mksimps: Proof.context -> thm -> thm list  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
99  | 
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: 
51611 
diff
changeset
 | 
100  | 
val set_mkcong: (Proof.context -> thm -> thm) -> Proof.context -> Proof.context  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
101  | 
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: 
51611 
diff
changeset
 | 
102  | 
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: 
51611 
diff
changeset
 | 
103  | 
val set_termless: (term * term -> bool) -> Proof.context -> Proof.context  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
104  | 
val set_subgoaler: (Proof.context -> int -> tactic) -> Proof.context -> Proof.context  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
105  | 
val solver: Proof.context -> solver -> int -> tactic  | 
| 
39163
 
4d701c0388c3
more explicit indication of Config.raw options, which are only needed for bootstrapping Pure;
 
wenzelm 
parents: 
39116 
diff
changeset
 | 
106  | 
val simp_depth_limit_raw: Config.raw  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
107  | 
val default_mk_sym: Proof.context -> thm -> thm option  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
108  | 
val simproc_global_i: theory -> string -> term list ->  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
109  | 
(Proof.context -> term -> thm option) -> simproc  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
110  | 
val simproc_global: theory -> string -> string list ->  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
111  | 
(Proof.context -> term -> thm option) -> simproc  | 
| 41227 | 112  | 
val simp_trace_depth_limit_raw: Config.raw  | 
113  | 
val simp_trace_depth_limit_default: int Unsynchronized.ref  | 
|
114  | 
val simp_trace_default: bool Unsynchronized.ref  | 
|
115  | 
val simp_trace_raw: Config.raw  | 
|
116  | 
val simp_debug_raw: Config.raw  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
117  | 
val add_prems: thm list -> Proof.context -> Proof.context  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
118  | 
val set_reorient: (Proof.context -> term list -> term -> term -> bool) ->  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
119  | 
Proof.context -> Proof.context  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
120  | 
val set_solvers: solver list -> Proof.context -> Proof.context  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
121  | 
val rewrite_cterm: bool * bool * bool ->  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
122  | 
(Proof.context -> thm -> thm option) -> Proof.context -> conv  | 
| 16458 | 123  | 
val rewrite_term: theory -> thm list -> (term -> term option) list -> term -> term  | 
| 15023 | 124  | 
val rewrite_thm: bool * bool * bool ->  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
125  | 
(Proof.context -> thm -> thm option) -> Proof.context -> thm -> thm  | 
| 46465 | 126  | 
val generic_rewrite_goal_tac: bool * bool * bool ->  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
127  | 
(Proof.context -> tactic) -> Proof.context -> int -> tactic  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
54731 
diff
changeset
 | 
128  | 
val rewrite: Proof.context -> bool -> thm list -> conv  | 
| 10413 | 129  | 
end;  | 
130  | 
||
| 
41228
 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 
wenzelm 
parents: 
41227 
diff
changeset
 | 
131  | 
structure Raw_Simplifier: RAW_SIMPLIFIER =  | 
| 10413 | 132  | 
struct  | 
133  | 
||
| 15023 | 134  | 
(** datatype simpset **)  | 
135  | 
||
| 51590 | 136  | 
(* congruence rules *)  | 
137  | 
||
138  | 
type cong_name = bool * string;  | 
|
139  | 
||
140  | 
fun cong_name (Const (a, _)) = SOME (true, a)  | 
|
141  | 
| cong_name (Free (a, _)) = SOME (false, a)  | 
|
142  | 
| cong_name _ = NONE;  | 
|
143  | 
||
144  | 
||
| 15023 | 145  | 
(* rewrite rules *)  | 
| 10413 | 146  | 
|
| 
20546
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
147  | 
type rrule =  | 
| 
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
148  | 
 {thm: thm,         (*the rewrite rule*)
 | 
| 
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
149  | 
name: string, (*name of theorem from which rewrite rule was extracted*)  | 
| 
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
150  | 
lhs: term, (*the left-hand side*)  | 
| 
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
151  | 
elhs: cterm, (*the etac-contracted lhs*)  | 
| 
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
152  | 
extra: bool, (*extra variables outside of elhs*)  | 
| 
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
153  | 
fo: bool, (*use first-order matching*)  | 
| 
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
154  | 
perm: bool}; (*the rewrite rule is permutative*)  | 
| 15023 | 155  | 
|
| 
20546
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
156  | 
(*  | 
| 12603 | 157  | 
Remarks:  | 
| 10413 | 158  | 
- elhs is used for matching,  | 
| 15023 | 159  | 
lhs only for preservation of bound variable names;  | 
| 10413 | 160  | 
- fo is set iff  | 
161  | 
either elhs is first-order (no Var is applied),  | 
|
| 15023 | 162  | 
in which case fo-matching is complete,  | 
| 10413 | 163  | 
or elhs is not a pattern,  | 
| 
20546
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
164  | 
in which case there is nothing better to do;  | 
| 
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
165  | 
*)  | 
| 10413 | 166  | 
|
167  | 
fun eq_rrule ({thm = thm1, ...}: rrule, {thm = thm2, ...}: rrule) =
 | 
|
| 
22360
 
26ead7ed4f4b
moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
 
wenzelm 
parents: 
22254 
diff
changeset
 | 
168  | 
Thm.eq_thm_prop (thm1, thm2);  | 
| 15023 | 169  | 
|
| 
20546
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
170  | 
(* FIXME: it seems that the conditions on extra variables are too liberal if  | 
| 
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
171  | 
prems are nonempty: does solving the prems really guarantee instantiation of  | 
| 
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
172  | 
all its Vars? Better: a dynamic check each time a rule is applied.  | 
| 
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
173  | 
*)  | 
| 
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
174  | 
fun rewrite_rule_extra_vars prems elhs erhs =  | 
| 
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
175  | 
let  | 
| 
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
176  | 
val elhss = elhs :: prems;  | 
| 
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
177  | 
val tvars = fold Term.add_tvars elhss [];  | 
| 
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
178  | 
val vars = fold Term.add_vars elhss [];  | 
| 
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
179  | 
in  | 
| 
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
180  | 
erhs |> Term.exists_type (Term.exists_subtype  | 
| 
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
181  | 
(fn TVar v => not (member (op =) tvars v) | _ => false)) orelse  | 
| 
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
182  | 
erhs |> Term.exists_subterm  | 
| 
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
183  | 
(fn Var v => not (member (op =) vars v) | _ => false)  | 
| 
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
184  | 
end;  | 
| 
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
185  | 
|
| 
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
186  | 
fun rrule_extra_vars elhs thm =  | 
| 
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
187  | 
rewrite_rule_extra_vars [] (term_of elhs) (Thm.full_prop_of thm);  | 
| 
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
188  | 
|
| 15023 | 189  | 
fun mk_rrule2 {thm, name, lhs, elhs, perm} =
 | 
190  | 
let  | 
|
| 
20546
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
191  | 
val t = term_of elhs;  | 
| 
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
192  | 
val fo = Pattern.first_order t orelse not (Pattern.pattern t);  | 
| 
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
193  | 
val extra = rrule_extra_vars elhs thm;  | 
| 
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
194  | 
  in {thm = thm, name = name, lhs = lhs, elhs = elhs, extra = extra, fo = fo, perm = perm} end;
 | 
| 10413 | 195  | 
|
| 15023 | 196  | 
(*simple test for looping rewrite rules and stupid orientations*)  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
197  | 
fun default_reorient ctxt prems lhs rhs =  | 
| 15023 | 198  | 
rewrite_rule_extra_vars prems lhs rhs  | 
199  | 
orelse  | 
|
200  | 
is_Var (head_of lhs)  | 
|
201  | 
orelse  | 
|
| 16305 | 202  | 
(* turns t = x around, which causes a headache if x is a local variable -  | 
203  | 
usually it is very useful :-(  | 
|
204  | 
is_Free rhs andalso not(is_Free lhs) andalso not(Logic.occs(rhs,lhs))  | 
|
205  | 
andalso not(exists_subterm is_Var lhs)  | 
|
206  | 
orelse  | 
|
207  | 
*)  | 
|
| 16842 | 208  | 
exists (fn t => Logic.occs (lhs, t)) (rhs :: prems)  | 
| 15023 | 209  | 
orelse  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
210  | 
null prems andalso Pattern.matches (Proof_Context.theory_of ctxt) (lhs, rhs)  | 
| 10413 | 211  | 
(*the condition "null prems" is necessary because conditional rewrites  | 
212  | 
with extra variables in the conditions may terminate although  | 
|
| 15023 | 213  | 
the rhs is an instance of the lhs; example: ?m < ?n ==> f(?n) == f(?m)*)  | 
214  | 
orelse  | 
|
215  | 
is_Const lhs andalso not (is_Const rhs);  | 
|
| 10413 | 216  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
217  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
218  | 
(* simplification procedures *)  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
219  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
220  | 
datatype proc =  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
221  | 
Proc of  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
222  | 
   {name: string,
 | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
223  | 
lhs: cterm,  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
224  | 
proc: Proof.context -> cterm -> thm option,  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
225  | 
id: stamp * thm list};  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
226  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
227  | 
fun eq_procid ((s1: stamp, ths1: thm list), (s2, ths2)) =  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
228  | 
s1 = s2 andalso eq_list Thm.eq_thm (ths1, ths2);  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
229  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
230  | 
fun eq_proc (Proc {id = id1, ...}, Proc {id = id2, ...}) = eq_procid (id1, id2);
 | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
231  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
232  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
233  | 
(* solvers *)  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
234  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
235  | 
datatype solver =  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
236  | 
Solver of  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
237  | 
   {name: string,
 | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
238  | 
solver: Proof.context -> int -> tactic,  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
239  | 
id: stamp};  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
240  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
241  | 
fun mk_solver name solver = Solver {name = name, solver = solver, id = stamp ()};
 | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
242  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
243  | 
fun solver_name (Solver {name, ...}) = name;
 | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
244  | 
fun solver ctxt (Solver {solver = tac, ...}) = tac ctxt;
 | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
245  | 
fun eq_solver (Solver {id = id1, ...}, Solver {id = id2, ...}) = (id1 = id2);
 | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
246  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
247  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
248  | 
(* simplification sets *)  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
249  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
250  | 
(*A simpset contains data required during conversion:  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
251  | 
rules: discrimination net of rewrite rules;  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
252  | 
prems: current premises;  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
253  | 
depth: simp_depth and exceeded flag;  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
254  | 
congs: association list of congruence rules and  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
255  | 
a list of `weak' congruence constants.  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
256  | 
A congruence is `weak' if it avoids normalization of some argument.  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
257  | 
procs: discrimination net of simplification procedures  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
258  | 
(functions that prove rewrite rules on the fly);  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
259  | 
mk_rews:  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
260  | 
mk: turn simplification thms into rewrite rules;  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
261  | 
mk_cong: prepare congruence rules;  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
262  | 
mk_sym: turn == around;  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
263  | 
mk_eq_True: turn P into P == True;  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
264  | 
termless: relation for ordered rewriting;*)  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
265  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
266  | 
datatype simpset =  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
267  | 
Simpset of  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
268  | 
   {rules: rrule Net.net,
 | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
269  | 
prems: thm list,  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
270  | 
depth: int * bool Unsynchronized.ref} *  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
271  | 
   {congs: (cong_name * thm) list * cong_name list,
 | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
272  | 
procs: proc Net.net,  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
273  | 
mk_rews:  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
274  | 
     {mk: Proof.context -> thm -> thm list,
 | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
275  | 
mk_cong: Proof.context -> thm -> thm,  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
276  | 
mk_sym: Proof.context -> thm -> thm option,  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
277  | 
mk_eq_True: Proof.context -> thm -> thm option,  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
278  | 
reorient: Proof.context -> term list -> term -> term -> bool},  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
279  | 
termless: term * term -> bool,  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
280  | 
subgoal_tac: Proof.context -> int -> tactic,  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
281  | 
loop_tacs: (string * (Proof.context -> int -> tactic)) list,  | 
| 
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: 
54729 
diff
changeset
 | 
282  | 
solvers: solver list * solver list};  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
283  | 
|
| 54728 | 284  | 
fun internal_ss (Simpset (_, ss2)) = ss2;  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
285  | 
|
| 
55014
 
a93f496f6c30
general notion of auxiliary bounds within context;
 
wenzelm 
parents: 
55000 
diff
changeset
 | 
286  | 
fun make_ss1 (rules, prems, depth) = {rules = rules, prems = prems, depth = depth};
 | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
287  | 
|
| 
55014
 
a93f496f6c30
general notion of auxiliary bounds within context;
 
wenzelm 
parents: 
55000 
diff
changeset
 | 
288  | 
fun map_ss1 f {rules, prems, depth} = make_ss1 (f (rules, prems, depth));
 | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
289  | 
|
| 
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: 
54729 
diff
changeset
 | 
290  | 
fun make_ss2 (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) =  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
291  | 
  {congs = congs, procs = procs, mk_rews = mk_rews, termless = termless,
 | 
| 
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: 
54729 
diff
changeset
 | 
292  | 
subgoal_tac = subgoal_tac, loop_tacs = loop_tacs, solvers = solvers};  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
293  | 
|
| 
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: 
54729 
diff
changeset
 | 
294  | 
fun map_ss2 f {congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers} =
 | 
| 
 
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: 
54729 
diff
changeset
 | 
295  | 
make_ss2 (f (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers));  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
296  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
297  | 
fun make_simpset (args1, args2) = Simpset (make_ss1 args1, make_ss2 args2);  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
298  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
299  | 
fun dest_ss (Simpset ({rules, ...}, {congs, procs, loop_tacs, solvers, ...})) =
 | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
300  | 
 {simps = Net.entries rules
 | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
301  | 
    |> map (fn {name, thm, ...} => (name, thm)),
 | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
302  | 
procs = Net.entries procs  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
303  | 
    |> map (fn Proc {name, lhs, id, ...} => ((name, lhs), id))
 | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
304  | 
|> partition_eq (eq_snd eq_procid)  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
305  | 
|> map (fn ps => (fst (fst (hd ps)), map (snd o fst) ps)),  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
306  | 
congs = #1 congs,  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
307  | 
weak_congs = #2 congs,  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
308  | 
loopers = map fst loop_tacs,  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
309  | 
unsafe_solvers = map solver_name (#1 solvers),  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
310  | 
safe_solvers = map solver_name (#2 solvers)};  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
311  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
312  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
313  | 
(* empty *)  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
314  | 
|
| 
55014
 
a93f496f6c30
general notion of auxiliary bounds within context;
 
wenzelm 
parents: 
55000 
diff
changeset
 | 
315  | 
fun init_ss depth mk_rews termless subgoal_tac solvers =  | 
| 
 
a93f496f6c30
general notion of auxiliary bounds within context;
 
wenzelm 
parents: 
55000 
diff
changeset
 | 
316  | 
make_simpset ((Net.empty, [], depth),  | 
| 
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: 
54729 
diff
changeset
 | 
317  | 
(([], []), Net.empty, mk_rews, termless, subgoal_tac, [], solvers));  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
318  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
319  | 
fun default_mk_sym _ th = SOME (th RS Drule.symmetric_thm);  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
320  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
321  | 
val empty_ss =  | 
| 
55014
 
a93f496f6c30
general notion of auxiliary bounds within context;
 
wenzelm 
parents: 
55000 
diff
changeset
 | 
322  | 
init_ss (0, Unsynchronized.ref false)  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
323  | 
    {mk = fn _ => fn th => if can Logic.dest_equals (Thm.concl_of th) then [th] else [],
 | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
324  | 
mk_cong = K I,  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
325  | 
mk_sym = default_mk_sym,  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
326  | 
mk_eq_True = K (K NONE),  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
327  | 
reorient = default_reorient}  | 
| 
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: 
54729 
diff
changeset
 | 
328  | 
Term_Ord.termless (K (K no_tac)) ([], []);  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
329  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
330  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
331  | 
(* merge *) (*NOTE: ignores some fields of 2nd simpset*)  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
332  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
333  | 
fun merge_ss (ss1, ss2) =  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
334  | 
if pointer_eq (ss1, ss2) then ss1  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
335  | 
else  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
336  | 
let  | 
| 
55014
 
a93f496f6c30
general notion of auxiliary bounds within context;
 
wenzelm 
parents: 
55000 
diff
changeset
 | 
337  | 
      val Simpset ({rules = rules1, prems = prems1, depth = depth1},
 | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
338  | 
       {congs = (congs1, weak1), procs = procs1, mk_rews, termless, subgoal_tac,
 | 
| 
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: 
54729 
diff
changeset
 | 
339  | 
loop_tacs = loop_tacs1, solvers = (unsafe_solvers1, solvers1)}) = ss1;  | 
| 
55014
 
a93f496f6c30
general notion of auxiliary bounds within context;
 
wenzelm 
parents: 
55000 
diff
changeset
 | 
340  | 
      val Simpset ({rules = rules2, prems = prems2, depth = depth2},
 | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
341  | 
       {congs = (congs2, weak2), procs = procs2, mk_rews = _, termless = _, subgoal_tac = _,
 | 
| 
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: 
54729 
diff
changeset
 | 
342  | 
loop_tacs = loop_tacs2, solvers = (unsafe_solvers2, solvers2)}) = ss2;  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
343  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
344  | 
val rules' = Net.merge eq_rrule (rules1, rules2);  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
345  | 
val prems' = Thm.merge_thms (prems1, prems2);  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
346  | 
val depth' = if #1 depth1 < #1 depth2 then depth2 else depth1;  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
347  | 
val congs' = merge (Thm.eq_thm_prop o pairself #2) (congs1, congs2);  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
348  | 
val weak' = merge (op =) (weak1, weak2);  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
349  | 
val procs' = Net.merge eq_proc (procs1, procs2);  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
350  | 
val loop_tacs' = AList.merge (op =) (K true) (loop_tacs1, loop_tacs2);  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
351  | 
val unsafe_solvers' = merge eq_solver (unsafe_solvers1, unsafe_solvers2);  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
352  | 
val solvers' = merge eq_solver (solvers1, solvers2);  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
353  | 
in  | 
| 
55014
 
a93f496f6c30
general notion of auxiliary bounds within context;
 
wenzelm 
parents: 
55000 
diff
changeset
 | 
354  | 
make_simpset ((rules', prems', depth'), ((congs', weak'), procs',  | 
| 
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: 
54729 
diff
changeset
 | 
355  | 
mk_rews, termless, subgoal_tac, loop_tacs', (unsafe_solvers', solvers')))  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
356  | 
end;  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
357  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
358  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
359  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
360  | 
(** context data **)  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
361  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
362  | 
structure Simpset = Generic_Data  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
363  | 
(  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
364  | 
type T = simpset;  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
365  | 
val empty = empty_ss;  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
366  | 
val extend = I;  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
367  | 
val merge = merge_ss;  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
368  | 
);  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
369  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
370  | 
val simpset_of = Simpset.get o Context.Proof;  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
371  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
372  | 
fun map_simpset f = Context.proof_map (Simpset.map f);  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
373  | 
fun map_simpset1 f = map_simpset (fn Simpset (ss1, ss2) => Simpset (map_ss1 f ss1, ss2));  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
374  | 
fun map_simpset2 f = map_simpset (fn Simpset (ss1, ss2) => Simpset (ss1, map_ss2 f ss2));  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
375  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
376  | 
fun simpset_map ctxt f ss = ctxt |> map_simpset (K ss) |> f |> Context.Proof |> Simpset.get;  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
377  | 
|
| 
55377
 
d79c057c68f0
more elementary put_simpset: exchange the simplifier configuration outright, which is particularly relevant concerning cumulative depth, e.g. for Product_Type.split_beta in the subsequent example:
 
wenzelm 
parents: 
55316 
diff
changeset
 | 
378  | 
fun put_simpset ss = map_simpset (K ss);  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
379  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
380  | 
val empty_simpset = put_simpset empty_ss;  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
381  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
382  | 
fun map_theory_simpset f thy =  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
383  | 
let  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
384  | 
val ctxt' = f (Proof_Context.init_global thy);  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
385  | 
val thy' = Proof_Context.theory_of ctxt';  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
386  | 
in Context.theory_map (Simpset.map (K (simpset_of ctxt'))) thy' end;  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
387  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
388  | 
fun map_ss f = Context.mapping (map_theory_simpset f) f;  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
389  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
390  | 
val clear_simpset =  | 
| 
55014
 
a93f496f6c30
general notion of auxiliary bounds within context;
 
wenzelm 
parents: 
55000 
diff
changeset
 | 
391  | 
  map_simpset (fn Simpset ({depth, ...}, {mk_rews, termless, subgoal_tac, solvers, ...}) =>
 | 
| 
 
a93f496f6c30
general notion of auxiliary bounds within context;
 
wenzelm 
parents: 
55000 
diff
changeset
 | 
392  | 
init_ss depth mk_rews termless subgoal_tac solvers);  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
393  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
394  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
395  | 
(* simp depth *)  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
396  | 
|
| 56438 | 397  | 
val simp_depth_limit_raw = Config.declare ("simp_depth_limit", @{here}) (K (Config.Int 100));
 | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
398  | 
val simp_depth_limit = Config.int simp_depth_limit_raw;  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
399  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
400  | 
val simp_trace_depth_limit_default = Unsynchronized.ref 1;  | 
| 56438 | 401  | 
val simp_trace_depth_limit_raw =  | 
402  | 
  Config.declare ("simp_trace_depth_limit", @{here})
 | 
|
403  | 
(fn _ => Config.Int (! simp_trace_depth_limit_default));  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
404  | 
val simp_trace_depth_limit = Config.int simp_trace_depth_limit_raw;  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
405  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
406  | 
fun inc_simp_depth ctxt =  | 
| 
55014
 
a93f496f6c30
general notion of auxiliary bounds within context;
 
wenzelm 
parents: 
55000 
diff
changeset
 | 
407  | 
ctxt |> map_simpset1 (fn (rules, prems, (depth, exceeded)) =>  | 
| 
 
a93f496f6c30
general notion of auxiliary bounds within context;
 
wenzelm 
parents: 
55000 
diff
changeset
 | 
408  | 
(rules, prems,  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
409  | 
(depth + 1,  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
410  | 
if depth = Config.get ctxt simp_trace_depth_limit  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
411  | 
then Unsynchronized.ref false else exceeded)));  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
412  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
413  | 
fun simp_depth ctxt =  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
414  | 
  let val Simpset ({depth = (depth, _), ...}, _) = simpset_of ctxt
 | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
415  | 
in depth end;  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
416  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
417  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
418  | 
(* diagnostics *)  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
419  | 
|
| 54997 | 420  | 
exception SIMPLIFIER of string * thm list;  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
421  | 
|
| 56438 | 422  | 
val simp_debug_raw = Config.declare ("simp_debug", @{here}) (K (Config.Bool false));
 | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
423  | 
val simp_debug = Config.bool simp_debug_raw;  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
424  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
425  | 
val simp_trace_default = Unsynchronized.ref false;  | 
| 56438 | 426  | 
val simp_trace_raw =  | 
427  | 
  Config.declare ("simp_trace", @{here}) (fn _ => Config.Bool (! simp_trace_default));
 | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
428  | 
val simp_trace = Config.bool simp_trace_raw;  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
429  | 
|
| 
55028
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
430  | 
fun cond_warning ctxt msg =  | 
| 
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
431  | 
if Context_Position.is_visible ctxt then warning (msg ()) else ();  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
432  | 
|
| 55031 | 433  | 
fun cond_tracing' ctxt flag msg =  | 
| 
55028
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
434  | 
if Config.get ctxt flag then  | 
| 
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
435  | 
let  | 
| 
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
436  | 
      val Simpset ({depth = (depth, exceeded), ...}, _) = simpset_of ctxt;
 | 
| 
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
437  | 
val depth_limit = Config.get ctxt simp_trace_depth_limit;  | 
| 
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
438  | 
in  | 
| 
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
439  | 
if depth > depth_limit then  | 
| 
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
440  | 
if ! exceeded then () else (tracing "simp_trace_depth_limit exceeded!"; exceeded := true)  | 
| 
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
441  | 
else (tracing (enclose "[" "]" (string_of_int depth) ^ msg ()); exceeded := false)  | 
| 
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
442  | 
end  | 
| 
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
443  | 
else ();  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
444  | 
|
| 55031 | 445  | 
fun cond_tracing ctxt = cond_tracing' ctxt simp_trace;  | 
446  | 
||
| 
55028
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
447  | 
fun print_term ctxt s t =  | 
| 
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
448  | 
s ^ "\n" ^ Syntax.string_of_term ctxt t;  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
449  | 
|
| 
55028
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
450  | 
fun print_thm ctxt s (name, th) =  | 
| 
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
451  | 
print_term ctxt (if name = "" then s else s ^ " " ^ quote name ^ ":") (Thm.full_prop_of th);  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
452  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
453  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
454  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
455  | 
(** simpset operations **)  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
456  | 
|
| 
55014
 
a93f496f6c30
general notion of auxiliary bounds within context;
 
wenzelm 
parents: 
55000 
diff
changeset
 | 
457  | 
(* prems *)  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
458  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
459  | 
fun prems_of ctxt =  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
460  | 
  let val Simpset ({prems, ...}, _) = simpset_of ctxt in prems end;
 | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
461  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
462  | 
fun add_prems ths =  | 
| 
55014
 
a93f496f6c30
general notion of auxiliary bounds within context;
 
wenzelm 
parents: 
55000 
diff
changeset
 | 
463  | 
map_simpset1 (fn (rules, prems, depth) => (rules, ths @ prems, depth));  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
464  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
465  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
466  | 
(* maintain simp rules *)  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
467  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
468  | 
fun del_rrule (rrule as {thm, elhs, ...}) ctxt =
 | 
| 
55014
 
a93f496f6c30
general notion of auxiliary bounds within context;
 
wenzelm 
parents: 
55000 
diff
changeset
 | 
469  | 
ctxt |> map_simpset1 (fn (rules, prems, depth) =>  | 
| 
 
a93f496f6c30
general notion of auxiliary bounds within context;
 
wenzelm 
parents: 
55000 
diff
changeset
 | 
470  | 
(Net.delete_term eq_rrule (term_of elhs, rrule) rules, prems, depth))  | 
| 
55028
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
471  | 
handle Net.DELETE =>  | 
| 
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
472  | 
    (cond_warning ctxt (fn () => print_thm ctxt "Rewrite rule not in simpset:" ("", thm)); ctxt);
 | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
473  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
474  | 
fun insert_rrule (rrule as {thm, name, ...}) ctxt =
 | 
| 55031 | 475  | 
(cond_tracing ctxt (fn () => print_thm ctxt "Adding rewrite rule" (name, thm));  | 
| 
55014
 
a93f496f6c30
general notion of auxiliary bounds within context;
 
wenzelm 
parents: 
55000 
diff
changeset
 | 
476  | 
ctxt |> map_simpset1 (fn (rules, prems, depth) =>  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
477  | 
let  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
478  | 
      val rrule2 as {elhs, ...} = mk_rrule2 rrule;
 | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
479  | 
val rules' = Net.insert_term eq_rrule (term_of elhs, rrule2) rules;  | 
| 
55014
 
a93f496f6c30
general notion of auxiliary bounds within context;
 
wenzelm 
parents: 
55000 
diff
changeset
 | 
480  | 
in (rules', prems, depth) end)  | 
| 
55028
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
481  | 
handle Net.INSERT =>  | 
| 
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
482  | 
    (cond_warning ctxt (fn () => print_thm ctxt "Ignoring duplicate rewrite rule:" ("", thm));
 | 
| 
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
483  | 
ctxt));  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
484  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
485  | 
local  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
486  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
487  | 
fun vperm (Var _, Var _) = true  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
488  | 
| vperm (Abs (_, _, s), Abs (_, _, t)) = vperm (s, t)  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
489  | 
| vperm (t1 $ t2, u1 $ u2) = vperm (t1, u1) andalso vperm (t2, u2)  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
490  | 
| vperm (t, u) = (t = u);  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
491  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
492  | 
fun var_perm (t, u) =  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
493  | 
vperm (t, u) andalso eq_set (op =) (Term.add_vars t [], Term.add_vars u []);  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
494  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
495  | 
in  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
496  | 
|
| 10413 | 497  | 
fun decomp_simp thm =  | 
| 15023 | 498  | 
let  | 
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
26424 
diff
changeset
 | 
499  | 
val prop = Thm.prop_of thm;  | 
| 15023 | 500  | 
val prems = Logic.strip_imp_prems prop;  | 
501  | 
val concl = Drule.strip_imp_concl (Thm.cprop_of thm);  | 
|
| 
22902
 
ac833b4bb7ee
moved some Drule operations to Thm (see more_thm.ML);
 
wenzelm 
parents: 
22892 
diff
changeset
 | 
502  | 
val (lhs, rhs) = Thm.dest_equals concl handle TERM _ =>  | 
| 54997 | 503  | 
      raise SIMPLIFIER ("Rewrite rule not a meta-equality", [thm]);
 | 
| 20579 | 504  | 
val elhs = Thm.dest_arg (Thm.cprop_of (Thm.eta_conversion lhs));  | 
| 18929 | 505  | 
val erhs = Envir.eta_contract (term_of rhs);  | 
| 15023 | 506  | 
val perm =  | 
507  | 
var_perm (term_of elhs, erhs) andalso  | 
|
508  | 
not (term_of elhs aconv erhs) andalso  | 
|
509  | 
not (is_Var (term_of elhs));  | 
|
| 52091 | 510  | 
in (prems, term_of lhs, elhs, term_of rhs, perm) end;  | 
| 10413 | 511  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
512  | 
end;  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
513  | 
|
| 12783 | 514  | 
fun decomp_simp' thm =  | 
| 52091 | 515  | 
let val (_, lhs, _, rhs, _) = decomp_simp thm in  | 
| 54997 | 516  | 
    if Thm.nprems_of thm > 0 then raise SIMPLIFIER ("Bad conditional rewrite rule", [thm])
 | 
| 
12979
 
4c76bce4ce39
decomp_simp': use lhs instead of elhs (preserves more bound variable names);
 
wenzelm 
parents: 
12783 
diff
changeset
 | 
517  | 
else (lhs, rhs)  | 
| 12783 | 518  | 
end;  | 
519  | 
||
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
520  | 
fun mk_eq_True ctxt (thm, name) =  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
521  | 
  let val Simpset (_, {mk_rews = {mk_eq_True, ...}, ...}) = simpset_of ctxt in
 | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
522  | 
(case mk_eq_True ctxt thm of  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
523  | 
NONE => []  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
524  | 
| SOME eq_True =>  | 
| 52091 | 525  | 
let val (_, lhs, elhs, _, _) = decomp_simp eq_True;  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
526  | 
        in [{thm = eq_True, name = name, lhs = lhs, elhs = elhs, perm = false}] end)
 | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
527  | 
end;  | 
| 10413 | 528  | 
|
| 15023 | 529  | 
(*create the rewrite rule and possibly also the eq_True variant,  | 
530  | 
in case there are extra vars on the rhs*)  | 
|
| 52082 | 531  | 
fun rrule_eq_True ctxt thm name lhs elhs rhs thm2 =  | 
| 15023 | 532  | 
  let val rrule = {thm = thm, name = name, lhs = lhs, elhs = elhs, perm = false} in
 | 
| 
20546
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
533  | 
if rewrite_rule_extra_vars [] lhs rhs then  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
534  | 
mk_eq_True ctxt (thm2, name) @ [rrule]  | 
| 
20546
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
535  | 
else [rrule]  | 
| 10413 | 536  | 
end;  | 
537  | 
||
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
538  | 
fun mk_rrule ctxt (thm, name) =  | 
| 52091 | 539  | 
let val (prems, lhs, elhs, rhs, perm) = decomp_simp thm in  | 
| 15023 | 540  | 
    if perm then [{thm = thm, name = name, lhs = lhs, elhs = elhs, perm = true}]
 | 
541  | 
else  | 
|
542  | 
(*weak test for loops*)  | 
|
543  | 
if rewrite_rule_extra_vars prems lhs rhs orelse is_Var (term_of elhs)  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
544  | 
then mk_eq_True ctxt (thm, name)  | 
| 52082 | 545  | 
else rrule_eq_True ctxt thm name lhs elhs rhs thm  | 
| 10413 | 546  | 
end;  | 
547  | 
||
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
548  | 
fun orient_rrule ctxt (thm, name) =  | 
| 18208 | 549  | 
let  | 
| 52091 | 550  | 
val (prems, lhs, elhs, rhs, perm) = decomp_simp thm;  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
551  | 
    val Simpset (_, {mk_rews = {reorient, mk_sym, ...}, ...}) = simpset_of ctxt;
 | 
| 18208 | 552  | 
in  | 
| 15023 | 553  | 
    if perm then [{thm = thm, name = name, lhs = lhs, elhs = elhs, perm = true}]
 | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
554  | 
else if reorient ctxt prems lhs rhs then  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
555  | 
if reorient ctxt prems rhs lhs  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
556  | 
then mk_eq_True ctxt (thm, name)  | 
| 15023 | 557  | 
else  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
558  | 
(case mk_sym ctxt thm of  | 
| 18208 | 559  | 
NONE => []  | 
560  | 
| SOME thm' =>  | 
|
| 52091 | 561  | 
let val (_, lhs', elhs', rhs', _) = decomp_simp thm'  | 
| 52082 | 562  | 
in rrule_eq_True ctxt thm' name lhs' elhs' rhs' thm end)  | 
563  | 
else rrule_eq_True ctxt thm name lhs elhs rhs thm  | 
|
| 10413 | 564  | 
end;  | 
565  | 
||
| 54982 | 566  | 
fun extract_rews ctxt thms =  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
567  | 
  let val Simpset (_, {mk_rews = {mk, ...}, ...}) = simpset_of ctxt
 | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
568  | 
in maps (fn thm => map (rpair (Thm.get_name_hint thm)) (mk ctxt thm)) thms end;  | 
| 10413 | 569  | 
|
| 54982 | 570  | 
fun extract_safe_rrules ctxt thm =  | 
571  | 
maps (orient_rrule ctxt) (extract_rews ctxt [thm]);  | 
|
| 10413 | 572  | 
|
| 
55316
 
885500f4aa6a
interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
 
Lars Hupel <lars.hupel@mytum.de> 
parents: 
55032 
diff
changeset
 | 
573  | 
fun mk_rrules ctxt thms =  | 
| 
 
885500f4aa6a
interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
 
Lars Hupel <lars.hupel@mytum.de> 
parents: 
55032 
diff
changeset
 | 
574  | 
let  | 
| 
 
885500f4aa6a
interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
 
Lars Hupel <lars.hupel@mytum.de> 
parents: 
55032 
diff
changeset
 | 
575  | 
val rews = extract_rews ctxt thms  | 
| 
 
885500f4aa6a
interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
 
Lars Hupel <lars.hupel@mytum.de> 
parents: 
55032 
diff
changeset
 | 
576  | 
val raw_rrules = flat (map (mk_rrule ctxt) rews)  | 
| 
 
885500f4aa6a
interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
 
Lars Hupel <lars.hupel@mytum.de> 
parents: 
55032 
diff
changeset
 | 
577  | 
in map mk_rrule2 raw_rrules end  | 
| 
 
885500f4aa6a
interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
 
Lars Hupel <lars.hupel@mytum.de> 
parents: 
55032 
diff
changeset
 | 
578  | 
|
| 10413 | 579  | 
|
| 
20028
 
b9752164ad92
add/del_simps: warning for inactive simpset (no context);
 
wenzelm 
parents: 
19798 
diff
changeset
 | 
580  | 
(* add/del rules explicitly *)  | 
| 10413 | 581  | 
|
| 54982 | 582  | 
fun comb_simps ctxt comb mk_rrule thms =  | 
| 
20028
 
b9752164ad92
add/del_simps: warning for inactive simpset (no context);
 
wenzelm 
parents: 
19798 
diff
changeset
 | 
583  | 
let  | 
| 54982 | 584  | 
val rews = extract_rews ctxt thms;  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
585  | 
in fold (fold comb o mk_rrule) rews ctxt end;  | 
| 10413 | 586  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
587  | 
fun ctxt addsimps thms =  | 
| 54982 | 588  | 
comb_simps ctxt insert_rrule (mk_rrule ctxt) thms;  | 
| 10413 | 589  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
590  | 
fun ctxt delsimps thms =  | 
| 54982 | 591  | 
comb_simps ctxt del_rrule (map mk_rrule2 o mk_rrule ctxt) thms;  | 
| 15023 | 592  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
593  | 
fun add_simp thm ctxt = ctxt addsimps [thm];  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
594  | 
fun del_simp thm ctxt = ctxt delsimps [thm];  | 
| 15023 | 595  | 
|
596  | 
(* congs *)  | 
|
| 10413 | 597  | 
|
| 15023 | 598  | 
local  | 
599  | 
||
600  | 
fun is_full_cong_prems [] [] = true  | 
|
601  | 
| is_full_cong_prems [] _ = false  | 
|
602  | 
| is_full_cong_prems (p :: prems) varpairs =  | 
|
603  | 
(case Logic.strip_assums_concl p of  | 
|
| 56245 | 604  | 
        Const ("Pure.eq", _) $ lhs $ rhs =>
 | 
| 15023 | 605  | 
let val (x, xs) = strip_comb lhs and (y, ys) = strip_comb rhs in  | 
606  | 
is_Var x andalso forall is_Bound xs andalso  | 
|
| 20972 | 607  | 
not (has_duplicates (op =) xs) andalso xs = ys andalso  | 
| 20671 | 608  | 
member (op =) varpairs (x, y) andalso  | 
| 19303 | 609  | 
is_full_cong_prems prems (remove (op =) (x, y) varpairs)  | 
| 15023 | 610  | 
end  | 
611  | 
| _ => false);  | 
|
612  | 
||
613  | 
fun is_full_cong thm =  | 
|
| 10413 | 614  | 
let  | 
| 43597 | 615  | 
val prems = Thm.prems_of thm and concl = Thm.concl_of thm;  | 
| 15023 | 616  | 
val (lhs, rhs) = Logic.dest_equals concl;  | 
617  | 
val (f, xs) = strip_comb lhs and (g, ys) = strip_comb rhs;  | 
|
| 10413 | 618  | 
in  | 
| 20972 | 619  | 
f = g andalso not (has_duplicates (op =) (xs @ ys)) andalso length xs = length ys andalso  | 
| 15023 | 620  | 
is_full_cong_prems prems (xs ~~ ys)  | 
| 10413 | 621  | 
end;  | 
622  | 
||
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
623  | 
fun mk_cong ctxt =  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
624  | 
  let val Simpset (_, {mk_rews = {mk_cong = f, ...}, ...}) = simpset_of ctxt
 | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
625  | 
in f ctxt end;  | 
| 
45620
 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 
wenzelm 
parents: 
45405 
diff
changeset
 | 
626  | 
|
| 
 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 
wenzelm 
parents: 
45405 
diff
changeset
 | 
627  | 
in  | 
| 
 
f2a587696afb
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 
wenzelm 
parents: 
45405 
diff
changeset
 | 
628  | 
|
| 
54729
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
629  | 
fun add_eqcong thm ctxt = ctxt |> map_simpset2  | 
| 
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: 
54729 
diff
changeset
 | 
630  | 
(fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) =>  | 
| 15023 | 631  | 
let  | 
| 45621 | 632  | 
val (lhs, _) = Logic.dest_equals (Thm.concl_of thm)  | 
| 54997 | 633  | 
        handle TERM _ => raise SIMPLIFIER ("Congruence not a meta-equality", [thm]);
 | 
| 18929 | 634  | 
(*val lhs = Envir.eta_contract lhs;*)  | 
| 45621 | 635  | 
val a = the (cong_name (head_of lhs)) handle Option.Option =>  | 
| 54997 | 636  | 
        raise SIMPLIFIER ("Congruence must start with a constant or free variable", [thm]);
 | 
| 
22221
 
8a8aa6114a89
changed cong alist - now using AList operations instead of overwrite_warn
 
haftmann 
parents: 
22008 
diff
changeset
 | 
637  | 
val (xs, weak) = congs;  | 
| 
38834
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
638  | 
val _ =  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
639  | 
if AList.defined (op =) xs a then  | 
| 
55028
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
640  | 
cond_warning ctxt (fn () => "Overwriting congruence rule for " ^ quote (#2 a))  | 
| 
22221
 
8a8aa6114a89
changed cong alist - now using AList operations instead of overwrite_warn
 
haftmann 
parents: 
22008 
diff
changeset
 | 
641  | 
else ();  | 
| 
30908
 
7ccf4a3d764c
replace type cong = {thm : thm, lhs : term} by plain thm -- the other component has been unused for a long time.
 
krauss 
parents: 
30552 
diff
changeset
 | 
642  | 
val xs' = AList.update (op =) (a, thm) xs;  | 
| 
22221
 
8a8aa6114a89
changed cong alist - now using AList operations instead of overwrite_warn
 
haftmann 
parents: 
22008 
diff
changeset
 | 
643  | 
val weak' = if is_full_cong thm then weak else a :: weak;  | 
| 
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: 
54729 
diff
changeset
 | 
644  | 
in ((xs', weak'), procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) end);  | 
| 10413 | 645  | 
|
| 
54729
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
646  | 
fun del_eqcong thm ctxt = ctxt |> map_simpset2  | 
| 
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: 
54729 
diff
changeset
 | 
647  | 
(fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) =>  | 
| 15023 | 648  | 
let  | 
| 45621 | 649  | 
val (lhs, _) = Logic.dest_equals (Thm.concl_of thm)  | 
| 54997 | 650  | 
        handle TERM _ => raise SIMPLIFIER ("Congruence not a meta-equality", [thm]);
 | 
| 18929 | 651  | 
(*val lhs = Envir.eta_contract lhs;*)  | 
| 20057 | 652  | 
val a = the (cong_name (head_of lhs)) handle Option.Option =>  | 
| 54997 | 653  | 
        raise SIMPLIFIER ("Congruence must start with a constant", [thm]);
 | 
| 
22221
 
8a8aa6114a89
changed cong alist - now using AList operations instead of overwrite_warn
 
haftmann 
parents: 
22008 
diff
changeset
 | 
654  | 
val (xs, _) = congs;  | 
| 51590 | 655  | 
val xs' = filter_out (fn (x : cong_name, _) => x = a) xs;  | 
| 
30908
 
7ccf4a3d764c
replace type cong = {thm : thm, lhs : term} by plain thm -- the other component has been unused for a long time.
 
krauss 
parents: 
30552 
diff
changeset
 | 
656  | 
val weak' = xs' |> map_filter (fn (a, thm) =>  | 
| 15531 | 657  | 
if is_full_cong thm then NONE else SOME a);  | 
| 
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: 
54729 
diff
changeset
 | 
658  | 
in ((xs', weak'), procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) end);  | 
| 10413 | 659  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
660  | 
fun add_cong thm ctxt = add_eqcong (mk_cong ctxt thm) ctxt;  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
661  | 
fun del_cong thm ctxt = del_eqcong (mk_cong ctxt thm) ctxt;  | 
| 15023 | 662  | 
|
663  | 
end;  | 
|
| 10413 | 664  | 
|
665  | 
||
| 15023 | 666  | 
(* simprocs *)  | 
667  | 
||
| 22234 | 668  | 
datatype simproc =  | 
669  | 
Simproc of  | 
|
670  | 
    {name: string,
 | 
|
671  | 
lhss: cterm list,  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
672  | 
proc: morphism -> Proof.context -> cterm -> thm option,  | 
| 22234 | 673  | 
id: stamp * thm list};  | 
674  | 
||
675  | 
fun eq_simproc (Simproc {id = id1, ...}, Simproc {id = id2, ...}) = eq_procid (id1, id2);
 | 
|
| 22008 | 676  | 
|
| 45290 | 677  | 
fun transform_simproc phi (Simproc {name, lhss, proc, id = (s, ths)}) =
 | 
| 22234 | 678  | 
Simproc  | 
679  | 
   {name = name,
 | 
|
680  | 
lhss = map (Morphism.cterm phi) lhss,  | 
|
| 22669 | 681  | 
proc = Morphism.transform phi proc,  | 
| 22234 | 682  | 
id = (s, Morphism.fact phi ths)};  | 
683  | 
||
684  | 
fun make_simproc {name, lhss, proc, identifier} =
 | 
|
685  | 
  Simproc {name = name, lhss = lhss, proc = proc, id = (stamp (), identifier)};
 | 
|
| 22008 | 686  | 
|
687  | 
fun mk_simproc name lhss proc =  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
688  | 
  make_simproc {name = name, lhss = lhss, proc = fn _ => fn ctxt => fn ct =>
 | 
| 
55028
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
689  | 
proc ctxt (term_of ct), identifier = []};  | 
| 22008 | 690  | 
|
| 
35845
 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 
wenzelm 
parents: 
35408 
diff
changeset
 | 
691  | 
(* FIXME avoid global thy and Logic.varify_global *)  | 
| 
38715
 
6513ea67d95d
renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
 
wenzelm 
parents: 
36944 
diff
changeset
 | 
692  | 
fun simproc_global_i thy name = mk_simproc name o map (Thm.cterm_of thy o Logic.varify_global);  | 
| 
 
6513ea67d95d
renamed Simplifier.simproc(_i) to Simplifier.simproc_global(_i) to emphasize that this is not the real thing;
 
wenzelm 
parents: 
36944 
diff
changeset
 | 
693  | 
fun simproc_global thy name = simproc_global_i thy name o map (Syntax.read_term_global thy);  | 
| 22008 | 694  | 
|
695  | 
||
| 15023 | 696  | 
local  | 
| 10413 | 697  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
698  | 
fun add_proc (proc as Proc {name, lhs, ...}) ctxt =
 | 
| 55031 | 699  | 
(cond_tracing ctxt (fn () =>  | 
| 
55028
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
700  | 
    print_term ctxt ("Adding simplification procedure " ^ quote name ^ " for") (term_of lhs));
 | 
| 
54729
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
701  | 
ctxt |> map_simpset2  | 
| 
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: 
54729 
diff
changeset
 | 
702  | 
(fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) =>  | 
| 
54729
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
703  | 
(congs, Net.insert_term eq_proc (term_of lhs, proc) procs,  | 
| 
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: 
54729 
diff
changeset
 | 
704  | 
mk_rews, termless, subgoal_tac, loop_tacs, solvers))  | 
| 15023 | 705  | 
handle Net.INSERT =>  | 
| 
55028
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
706  | 
(cond_warning ctxt (fn () => "Ignoring duplicate simplification procedure " ^ quote name);  | 
| 
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
707  | 
ctxt));  | 
| 10413 | 708  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
709  | 
fun del_proc (proc as Proc {name, lhs, ...}) ctxt =
 | 
| 
54729
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
710  | 
ctxt |> map_simpset2  | 
| 
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: 
54729 
diff
changeset
 | 
711  | 
(fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) =>  | 
| 
54729
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
712  | 
(congs, Net.delete_term eq_proc (term_of lhs, proc) procs,  | 
| 
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: 
54729 
diff
changeset
 | 
713  | 
mk_rews, termless, subgoal_tac, loop_tacs, solvers))  | 
| 15023 | 714  | 
handle Net.DELETE =>  | 
| 
55028
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
715  | 
(cond_warning ctxt (fn () => "Simplification procedure " ^ quote name ^ " not in simpset");  | 
| 
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
716  | 
ctxt);  | 
| 10413 | 717  | 
|
| 22234 | 718  | 
fun prep_procs (Simproc {name, lhss, proc, id}) =
 | 
| 22669 | 719  | 
  lhss |> map (fn lhs => Proc {name = name, lhs = lhs, proc = Morphism.form proc, id = id});
 | 
| 22234 | 720  | 
|
| 15023 | 721  | 
in  | 
| 10413 | 722  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
723  | 
fun ctxt addsimprocs ps = fold (fold add_proc o prep_procs) ps ctxt;  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
724  | 
fun ctxt delsimprocs ps = fold (fold del_proc o prep_procs) ps ctxt;  | 
| 10413 | 725  | 
|
| 15023 | 726  | 
end;  | 
| 10413 | 727  | 
|
728  | 
||
729  | 
(* mk_rews *)  | 
|
730  | 
||
| 15023 | 731  | 
local  | 
732  | 
||
| 
54729
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
733  | 
fun map_mk_rews f =  | 
| 
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: 
54729 
diff
changeset
 | 
734  | 
map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) =>  | 
| 
54729
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
735  | 
let  | 
| 
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
736  | 
      val {mk, mk_cong, mk_sym, mk_eq_True, reorient} = mk_rews;
 | 
| 
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
737  | 
val (mk', mk_cong', mk_sym', mk_eq_True', reorient') =  | 
| 
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
738  | 
f (mk, mk_cong, mk_sym, mk_eq_True, reorient);  | 
| 
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
739  | 
      val mk_rews' = {mk = mk', mk_cong = mk_cong', mk_sym = mk_sym', mk_eq_True = mk_eq_True',
 | 
| 
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
740  | 
reorient = reorient'};  | 
| 
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: 
54729 
diff
changeset
 | 
741  | 
in (congs, procs, mk_rews', termless, subgoal_tac, loop_tacs, solvers) end);  | 
| 15023 | 742  | 
|
743  | 
in  | 
|
| 10413 | 744  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
745  | 
fun mksimps ctxt =  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
746  | 
  let val Simpset (_, {mk_rews = {mk, ...}, ...}) = simpset_of ctxt
 | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
747  | 
in mk ctxt end;  | 
| 
30318
 
3d03190d2864
replaced archaic use of rep_ss by Simplifier.mksimps;
 
wenzelm 
parents: 
29269 
diff
changeset
 | 
748  | 
|
| 
45625
 
750c5a47400b
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 
wenzelm 
parents: 
45621 
diff
changeset
 | 
749  | 
fun set_mksimps mk = map_mk_rews (fn (_, mk_cong, mk_sym, mk_eq_True, reorient) =>  | 
| 18208 | 750  | 
(mk, mk_cong, mk_sym, mk_eq_True, reorient));  | 
| 15023 | 751  | 
|
| 
45625
 
750c5a47400b
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 
wenzelm 
parents: 
45621 
diff
changeset
 | 
752  | 
fun set_mkcong mk_cong = map_mk_rews (fn (mk, _, mk_sym, mk_eq_True, reorient) =>  | 
| 18208 | 753  | 
(mk, mk_cong, mk_sym, mk_eq_True, reorient));  | 
| 10413 | 754  | 
|
| 
45625
 
750c5a47400b
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 
wenzelm 
parents: 
45621 
diff
changeset
 | 
755  | 
fun set_mksym mk_sym = map_mk_rews (fn (mk, mk_cong, _, mk_eq_True, reorient) =>  | 
| 18208 | 756  | 
(mk, mk_cong, mk_sym, mk_eq_True, reorient));  | 
| 10413 | 757  | 
|
| 
45625
 
750c5a47400b
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 
wenzelm 
parents: 
45621 
diff
changeset
 | 
758  | 
fun set_mkeqTrue mk_eq_True = map_mk_rews (fn (mk, mk_cong, mk_sym, _, reorient) =>  | 
| 18208 | 759  | 
(mk, mk_cong, mk_sym, mk_eq_True, reorient));  | 
760  | 
||
761  | 
fun set_reorient reorient = map_mk_rews (fn (mk, mk_cong, mk_sym, mk_eq_True, _) =>  | 
|
762  | 
(mk, mk_cong, mk_sym, mk_eq_True, reorient));  | 
|
| 15023 | 763  | 
|
764  | 
end;  | 
|
765  | 
||
| 
14242
 
ec70653a02bf
Added access to the mk_rews field (and friends).
 
skalberg 
parents: 
14040 
diff
changeset
 | 
766  | 
|
| 10413 | 767  | 
(* termless *)  | 
768  | 
||
| 
45625
 
750c5a47400b
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 
wenzelm 
parents: 
45621 
diff
changeset
 | 
769  | 
fun set_termless termless =  | 
| 
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: 
54729 
diff
changeset
 | 
770  | 
map_simpset2 (fn (congs, procs, mk_rews, _, subgoal_tac, loop_tacs, solvers) =>  | 
| 
 
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: 
54729 
diff
changeset
 | 
771  | 
(congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers));  | 
| 
15006
 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 
skalberg 
parents: 
15001 
diff
changeset
 | 
772  | 
|
| 
 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 
skalberg 
parents: 
15001 
diff
changeset
 | 
773  | 
|
| 15023 | 774  | 
(* tactics *)  | 
| 
15006
 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 
skalberg 
parents: 
15001 
diff
changeset
 | 
775  | 
|
| 
45625
 
750c5a47400b
modernized some old-style infix operations, which were left over from the time of ML proof scripts;
 
wenzelm 
parents: 
45621 
diff
changeset
 | 
776  | 
fun set_subgoaler subgoal_tac =  | 
| 
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: 
54729 
diff
changeset
 | 
777  | 
map_simpset2 (fn (congs, procs, mk_rews, termless, _, loop_tacs, solvers) =>  | 
| 
 
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: 
54729 
diff
changeset
 | 
778  | 
(congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers));  | 
| 
15006
 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 
skalberg 
parents: 
15001 
diff
changeset
 | 
779  | 
|
| 52037 | 780  | 
fun ctxt setloop tac = ctxt |>  | 
| 
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: 
54729 
diff
changeset
 | 
781  | 
map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, _, solvers) =>  | 
| 
 
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: 
54729 
diff
changeset
 | 
782  | 
   (congs, procs, mk_rews, termless, subgoal_tac, [("", tac)], solvers));
 | 
| 
15006
 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 
skalberg 
parents: 
15001 
diff
changeset
 | 
783  | 
|
| 52037 | 784  | 
fun ctxt addloop (name, tac) = ctxt |>  | 
| 
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: 
54729 
diff
changeset
 | 
785  | 
map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) =>  | 
| 15023 | 786  | 
(congs, procs, mk_rews, termless, subgoal_tac,  | 
| 
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: 
54729 
diff
changeset
 | 
787  | 
AList.update (op =) (name, tac) loop_tacs, solvers));  | 
| 
15006
 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 
skalberg 
parents: 
15001 
diff
changeset
 | 
788  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
789  | 
fun ctxt delloop name = ctxt |>  | 
| 
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: 
54729 
diff
changeset
 | 
790  | 
map_simpset2 (fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, solvers) =>  | 
| 
21286
 
b5e7b80caa6a
introduces canonical AList functions for loop_tacs
 
haftmann 
parents: 
20972 
diff
changeset
 | 
791  | 
(congs, procs, mk_rews, termless, subgoal_tac,  | 
| 
38834
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
792  | 
(if AList.defined (op =) loop_tacs name then ()  | 
| 
55028
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
793  | 
else cond_warning ctxt (fn () => "No such looper in simpset: " ^ quote name);  | 
| 
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
794  | 
AList.delete (op =) name loop_tacs), solvers));  | 
| 
15006
 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 
skalberg 
parents: 
15001 
diff
changeset
 | 
795  | 
|
| 
54729
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
796  | 
fun ctxt setSSolver solver = ctxt |> map_simpset2  | 
| 
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: 
54729 
diff
changeset
 | 
797  | 
(fn (congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, (unsafe_solvers, _)) =>  | 
| 
 
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: 
54729 
diff
changeset
 | 
798  | 
(congs, procs, mk_rews, termless, subgoal_tac, loop_tacs, (unsafe_solvers, [solver])));  | 
| 
15006
 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 
skalberg 
parents: 
15001 
diff
changeset
 | 
799  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
800  | 
fun ctxt addSSolver solver = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, termless,  | 
| 
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: 
54729 
diff
changeset
 | 
801  | 
subgoal_tac, loop_tacs, (unsafe_solvers, solvers)) => (congs, procs, mk_rews, termless,  | 
| 
 
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: 
54729 
diff
changeset
 | 
802  | 
subgoal_tac, loop_tacs, (unsafe_solvers, insert eq_solver solver solvers)));  | 
| 
15006
 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 
skalberg 
parents: 
15001 
diff
changeset
 | 
803  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
804  | 
fun ctxt setSolver solver = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, termless,  | 
| 
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: 
54729 
diff
changeset
 | 
805  | 
subgoal_tac, loop_tacs, (_, solvers)) => (congs, procs, mk_rews, termless,  | 
| 
 
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: 
54729 
diff
changeset
 | 
806  | 
subgoal_tac, loop_tacs, ([solver], solvers)));  | 
| 
15006
 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 
skalberg 
parents: 
15001 
diff
changeset
 | 
807  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
808  | 
fun ctxt addSolver solver = ctxt |> map_simpset2 (fn (congs, procs, mk_rews, termless,  | 
| 
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: 
54729 
diff
changeset
 | 
809  | 
subgoal_tac, loop_tacs, (unsafe_solvers, solvers)) => (congs, procs, mk_rews, termless,  | 
| 
 
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: 
54729 
diff
changeset
 | 
810  | 
subgoal_tac, loop_tacs, (insert eq_solver solver unsafe_solvers, solvers)));  | 
| 
15006
 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 
skalberg 
parents: 
15001 
diff
changeset
 | 
811  | 
|
| 15023 | 812  | 
fun set_solvers solvers = map_simpset2 (fn (congs, procs, mk_rews, termless,  | 
| 
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: 
54729 
diff
changeset
 | 
813  | 
subgoal_tac, loop_tacs, _) => (congs, procs, mk_rews, termless,  | 
| 
 
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: 
54729 
diff
changeset
 | 
814  | 
subgoal_tac, loop_tacs, (solvers, solvers)));  | 
| 
54729
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
815  | 
|
| 
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
816  | 
|
| 
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
817  | 
(* trace operations *)  | 
| 
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
818  | 
|
| 
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: 
54729 
diff
changeset
 | 
819  | 
type trace_ops =  | 
| 
 
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: 
54729 
diff
changeset
 | 
820  | 
 {trace_invoke: {depth: int, term: term} -> Proof.context -> Proof.context,
 | 
| 
55316
 
885500f4aa6a
interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
 
Lars Hupel <lars.hupel@mytum.de> 
parents: 
55032 
diff
changeset
 | 
821  | 
  trace_apply: {unconditional: bool, term: term, thm: thm, rrule: rrule} ->
 | 
| 
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: 
54729 
diff
changeset
 | 
822  | 
Proof.context -> (Proof.context -> (thm * term) option) -> (thm * term) option};  | 
| 
54729
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
823  | 
|
| 
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: 
54729 
diff
changeset
 | 
824  | 
structure Trace_Ops = Theory_Data  | 
| 
 
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: 
54729 
diff
changeset
 | 
825  | 
(  | 
| 
 
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: 
54729 
diff
changeset
 | 
826  | 
type T = trace_ops;  | 
| 
 
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: 
54729 
diff
changeset
 | 
827  | 
val empty: T =  | 
| 
 
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: 
54729 
diff
changeset
 | 
828  | 
   {trace_invoke = fn _ => fn ctxt => ctxt,
 | 
| 
 
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: 
54729 
diff
changeset
 | 
829  | 
trace_apply = fn _ => fn ctxt => fn cont => cont ctxt};  | 
| 
 
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: 
54729 
diff
changeset
 | 
830  | 
val extend = I;  | 
| 
 
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: 
54729 
diff
changeset
 | 
831  | 
fun merge (trace_ops, _) = trace_ops;  | 
| 
 
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: 
54729 
diff
changeset
 | 
832  | 
);  | 
| 
 
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: 
54729 
diff
changeset
 | 
833  | 
|
| 
 
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: 
54729 
diff
changeset
 | 
834  | 
val set_trace_ops = Trace_Ops.put;  | 
| 
 
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: 
54729 
diff
changeset
 | 
835  | 
|
| 
 
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: 
54729 
diff
changeset
 | 
836  | 
val trace_ops = Trace_Ops.get o Proof_Context.theory_of;  | 
| 
54729
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
837  | 
fun trace_invoke args ctxt = #trace_invoke (trace_ops ctxt) args ctxt;  | 
| 
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
838  | 
fun trace_apply args ctxt = #trace_apply (trace_ops ctxt) args ctxt;  | 
| 
15006
 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 
skalberg 
parents: 
15001 
diff
changeset
 | 
839  | 
|
| 
 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 
skalberg 
parents: 
15001 
diff
changeset
 | 
840  | 
|
| 
 
107e4dfd3b96
Merging the meta-simplifier with the Provers-simplifier.  Next step:
 
skalberg 
parents: 
15001 
diff
changeset
 | 
841  | 
|
| 10413 | 842  | 
(** rewriting **)  | 
843  | 
||
844  | 
(*  | 
|
845  | 
Uses conversions, see:  | 
|
846  | 
L C Paulson, A higher-order implementation of rewriting,  | 
|
847  | 
Science of Computer Programming 3 (1983), pages 119-149.  | 
|
848  | 
*)  | 
|
849  | 
||
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
850  | 
fun check_conv ctxt msg thm thm' =  | 
| 10413 | 851  | 
let  | 
| 36944 | 852  | 
val thm'' = Thm.transitive thm thm' handle THM _ =>  | 
853  | 
Thm.transitive thm (Thm.transitive  | 
|
854  | 
(Thm.symmetric (Drule.beta_eta_conversion (Thm.lhs_of thm'))) thm')  | 
|
| 
55028
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
855  | 
val _ =  | 
| 55031 | 856  | 
      if msg then cond_tracing ctxt (fn () => print_thm ctxt "SUCCEEDED" ("", thm'))
 | 
| 
55028
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
857  | 
else ();  | 
| 
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
858  | 
in SOME thm'' end  | 
| 10413 | 859  | 
handle THM _ =>  | 
| 
26626
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
26424 
diff
changeset
 | 
860  | 
let  | 
| 
 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 
wenzelm 
parents: 
26424 
diff
changeset
 | 
861  | 
val _ $ _ $ prop0 = Thm.prop_of thm;  | 
| 
55028
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
862  | 
val _ =  | 
| 
55032
 
b49366215417
back to conditional tracing instead of noisy warning (see also 00e849f5b397): these incidents happen occasionally;
 
wenzelm 
parents: 
55031 
diff
changeset
 | 
863  | 
cond_tracing ctxt (fn () =>  | 
| 
 
b49366215417
back to conditional tracing instead of noisy warning (see also 00e849f5b397): these incidents happen occasionally;
 
wenzelm 
parents: 
55031 
diff
changeset
 | 
864  | 
          print_thm ctxt "Proved wrong theorem (bad subgoaler?)" ("", thm') ^ "\n" ^
 | 
| 
55028
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
865  | 
print_term ctxt "Should have proved:" prop0);  | 
| 
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
866  | 
in NONE end;  | 
| 10413 | 867  | 
|
868  | 
||
869  | 
(* mk_procrule *)  | 
|
870  | 
||
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
871  | 
fun mk_procrule ctxt thm =  | 
| 52091 | 872  | 
let val (prems, lhs, elhs, rhs, _) = decomp_simp thm in  | 
| 15023 | 873  | 
if rewrite_rule_extra_vars prems lhs rhs  | 
| 
55028
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
874  | 
    then (cond_warning ctxt (fn () => print_thm ctxt "Extra vars on rhs:" ("", thm)); [])
 | 
| 15023 | 875  | 
    else [mk_rrule2 {thm = thm, name = "", lhs = lhs, elhs = elhs, perm = false}]
 | 
| 10413 | 876  | 
end;  | 
877  | 
||
878  | 
||
| 15023 | 879  | 
(* rewritec: conversion to apply the meta simpset to a term *)  | 
| 10413 | 880  | 
|
| 15023 | 881  | 
(*Since the rewriting strategy is bottom-up, we avoid re-normalizing already  | 
882  | 
normalized terms by carrying around the rhs of the rewrite rule just  | 
|
883  | 
applied. This is called the `skeleton'. It is decomposed in parallel  | 
|
884  | 
with the term. Once a Var is encountered, the corresponding term is  | 
|
885  | 
already in normal form.  | 
|
886  | 
skel0 is a dummy skeleton that is to enforce complete normalization.*)  | 
|
887  | 
||
| 10413 | 888  | 
val skel0 = Bound 0;  | 
889  | 
||
| 15023 | 890  | 
(*Use rhs as skeleton only if the lhs does not contain unnormalized bits.  | 
891  | 
The latter may happen iff there are weak congruence rules for constants  | 
|
892  | 
in the lhs.*)  | 
|
| 10413 | 893  | 
|
| 15023 | 894  | 
fun uncond_skel ((_, weak), (lhs, rhs)) =  | 
895  | 
if null weak then rhs (*optimization*)  | 
|
| 
51591
 
e4aeb102ad70
amended uncond_skel to observe notion of cong_name properly -- may affect simplification with Free congs;
 
wenzelm 
parents: 
51590 
diff
changeset
 | 
896  | 
else if exists_subterm  | 
| 
 
e4aeb102ad70
amended uncond_skel to observe notion of cong_name properly -- may affect simplification with Free congs;
 
wenzelm 
parents: 
51590 
diff
changeset
 | 
897  | 
(fn Const (a, _) => member (op =) weak (true, a)  | 
| 
 
e4aeb102ad70
amended uncond_skel to observe notion of cong_name properly -- may affect simplification with Free congs;
 
wenzelm 
parents: 
51590 
diff
changeset
 | 
898  | 
| Free (a, _) => member (op =) weak (false, a)  | 
| 
 
e4aeb102ad70
amended uncond_skel to observe notion of cong_name properly -- may affect simplification with Free congs;
 
wenzelm 
parents: 
51590 
diff
changeset
 | 
899  | 
| _ => false) lhs then skel0  | 
| 15023 | 900  | 
else rhs;  | 
901  | 
||
902  | 
(*Behaves like unconditional rule if rhs does not contain vars not in the lhs.  | 
|
903  | 
Otherwise those vars may become instantiated with unnormalized terms  | 
|
904  | 
while the premises are solved.*)  | 
|
905  | 
||
| 32797 | 906  | 
fun cond_skel (args as (_, (lhs, rhs))) =  | 
| 33038 | 907  | 
if subset (op =) (Term.add_vars rhs [], Term.add_vars lhs []) then uncond_skel args  | 
| 10413 | 908  | 
else skel0;  | 
909  | 
||
910  | 
(*  | 
|
| 15023 | 911  | 
Rewriting -- we try in order:  | 
| 10413 | 912  | 
(1) beta reduction  | 
913  | 
(2) unconditional rewrite rules  | 
|
914  | 
(3) conditional rewrite rules  | 
|
915  | 
(4) simplification procedures  | 
|
916  | 
||
917  | 
IMPORTANT: rewrite rules must not introduce new Vars or TVars!  | 
|
918  | 
*)  | 
|
919  | 
||
| 52091 | 920  | 
fun rewritec (prover, maxt) ctxt t =  | 
| 10413 | 921  | 
let  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
922  | 
    val Simpset ({rules, ...}, {congs, procs, termless, ...}) = simpset_of ctxt;
 | 
| 10413 | 923  | 
val eta_thm = Thm.eta_conversion t;  | 
| 
22902
 
ac833b4bb7ee
moved some Drule operations to Thm (see more_thm.ML);
 
wenzelm 
parents: 
22892 
diff
changeset
 | 
924  | 
val eta_t' = Thm.rhs_of eta_thm;  | 
| 10413 | 925  | 
val eta_t = term_of eta_t';  | 
| 
55316
 
885500f4aa6a
interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
 
Lars Hupel <lars.hupel@mytum.de> 
parents: 
55032 
diff
changeset
 | 
926  | 
fun rew rrule =  | 
| 10413 | 927  | 
let  | 
| 
55316
 
885500f4aa6a
interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
 
Lars Hupel <lars.hupel@mytum.de> 
parents: 
55032 
diff
changeset
 | 
928  | 
        val {thm, name, lhs, elhs, extra, fo, perm} = rrule
 | 
| 32797 | 929  | 
val prop = Thm.prop_of thm;  | 
| 
20546
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
930  | 
val (rthm, elhs') =  | 
| 
 
8923deb735ad
rrule: maintain 'extra' field for rule that contain extra vars outside elhs;
 
wenzelm 
parents: 
20330 
diff
changeset
 | 
931  | 
if maxt = ~1 orelse not extra then (thm, elhs)  | 
| 
22902
 
ac833b4bb7ee
moved some Drule operations to Thm (see more_thm.ML);
 
wenzelm 
parents: 
22892 
diff
changeset
 | 
932  | 
else (Thm.incr_indexes (maxt + 1) thm, Thm.incr_indexes_cterm (maxt + 1) elhs);  | 
| 
 
ac833b4bb7ee
moved some Drule operations to Thm (see more_thm.ML);
 
wenzelm 
parents: 
22892 
diff
changeset
 | 
933  | 
val insts =  | 
| 
 
ac833b4bb7ee
moved some Drule operations to Thm (see more_thm.ML);
 
wenzelm 
parents: 
22892 
diff
changeset
 | 
934  | 
if fo then Thm.first_order_match (elhs', eta_t')  | 
| 
 
ac833b4bb7ee
moved some Drule operations to Thm (see more_thm.ML);
 
wenzelm 
parents: 
22892 
diff
changeset
 | 
935  | 
else Thm.match (elhs', eta_t');  | 
| 10413 | 936  | 
val thm' = Thm.instantiate insts (Thm.rename_boundvars lhs eta_t rthm);  | 
| 14643 | 937  | 
val prop' = Thm.prop_of thm';  | 
| 21576 | 938  | 
val unconditional = (Logic.count_prems prop' = 0);  | 
| 54725 | 939  | 
val (lhs', rhs') = Logic.dest_equals (Logic.strip_imp_concl prop');  | 
| 
55316
 
885500f4aa6a
interactive simplifier trace: new panel in Isabelle/jEdit to inspect and modify simplification state
 
Lars Hupel <lars.hupel@mytum.de> 
parents: 
55032 
diff
changeset
 | 
940  | 
        val trace_args = {unconditional = unconditional, term = eta_t, thm = thm', rrule = rrule};
 | 
| 10413 | 941  | 
in  | 
| 11295 | 942  | 
if perm andalso not (termless (rhs', lhs'))  | 
| 54725 | 943  | 
then  | 
| 55031 | 944  | 
(cond_tracing ctxt (fn () =>  | 
| 
55028
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
945  | 
print_thm ctxt "Cannot apply permutative rewrite rule" (name, thm) ^ "\n" ^  | 
| 
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
946  | 
            print_thm ctxt "Term does not become smaller:" ("", thm'));
 | 
| 54725 | 947  | 
NONE)  | 
948  | 
else  | 
|
| 55031 | 949  | 
(cond_tracing ctxt (fn () =>  | 
| 
55028
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
950  | 
print_thm ctxt "Applying instance of rewrite rule" (name, thm));  | 
| 54725 | 951  | 
if unconditional  | 
952  | 
then  | 
|
| 55031 | 953  | 
           (cond_tracing ctxt (fn () => print_thm ctxt "Rewriting:" ("", thm'));
 | 
| 
54729
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
954  | 
trace_apply trace_args ctxt (fn ctxt' =>  | 
| 
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
955  | 
let  | 
| 
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
956  | 
val lr = Logic.dest_equals prop;  | 
| 
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
957  | 
val SOME thm'' = check_conv ctxt' false eta_thm thm';  | 
| 
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
958  | 
in SOME (thm'', uncond_skel (congs, lr)) end))  | 
| 54725 | 959  | 
else  | 
| 55031 | 960  | 
           (cond_tracing ctxt (fn () => print_thm ctxt "Trying to rewrite:" ("", thm'));
 | 
| 54725 | 961  | 
if simp_depth ctxt > Config.get ctxt simp_depth_limit  | 
| 55031 | 962  | 
then (cond_tracing ctxt (fn () => "simp_depth_limit exceeded - giving up"); NONE)  | 
| 54725 | 963  | 
else  | 
| 
54729
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
964  | 
trace_apply trace_args ctxt (fn ctxt' =>  | 
| 
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
965  | 
(case prover ctxt' thm' of  | 
| 55031 | 966  | 
                  NONE => (cond_tracing ctxt' (fn () => print_thm ctxt' "FAILED" ("", thm')); NONE)
 | 
| 
54729
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
967  | 
| SOME thm2 =>  | 
| 
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
968  | 
(case check_conv ctxt' true eta_thm thm2 of  | 
| 
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
969  | 
NONE => NONE  | 
| 
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
970  | 
| SOME thm2' =>  | 
| 
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
971  | 
let  | 
| 
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
972  | 
val concl = Logic.strip_imp_concl prop;  | 
| 
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
973  | 
val lr = Logic.dest_equals concl;  | 
| 
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
974  | 
in SOME (thm2', cond_skel (congs, lr)) end)))))  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
975  | 
end;  | 
| 10413 | 976  | 
|
| 15531 | 977  | 
fun rews [] = NONE  | 
| 10413 | 978  | 
| rews (rrule :: rrules) =  | 
| 15531 | 979  | 
let val opt = rew rrule handle Pattern.MATCH => NONE  | 
| 54725 | 980  | 
in (case opt of NONE => rews rrules | some => some) end;  | 
| 10413 | 981  | 
|
| 
38834
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
982  | 
fun sort_rrules rrs =  | 
| 
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
983  | 
let  | 
| 
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
984  | 
        fun is_simple ({thm, ...}: rrule) =
 | 
| 
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
985  | 
(case Thm.prop_of thm of  | 
| 56245 | 986  | 
            Const ("Pure.eq", _) $ _ $ _ => true
 | 
| 
38834
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
987  | 
| _ => false);  | 
| 
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
988  | 
fun sort [] (re1, re2) = re1 @ re2  | 
| 
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
989  | 
| sort (rr :: rrs) (re1, re2) =  | 
| 
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
990  | 
if is_simple rr  | 
| 
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
991  | 
then sort rrs (rr :: re1, re2)  | 
| 
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
992  | 
else sort rrs (re1, rr :: re2);  | 
| 
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
993  | 
in sort rrs ([], []) end;  | 
| 10413 | 994  | 
|
| 15531 | 995  | 
fun proc_rews [] = NONE  | 
| 15023 | 996  | 
      | proc_rews (Proc {name, proc, lhs, ...} :: ps) =
 | 
| 
55028
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
997  | 
if Pattern.matches (Proof_Context.theory_of ctxt) (term_of lhs, term_of t) then  | 
| 55031 | 998  | 
(cond_tracing' ctxt simp_debug (fn () =>  | 
| 
55028
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
999  | 
              print_term ctxt ("Trying procedure " ^ quote name ^ " on:") eta_t);
 | 
| 54725 | 1000  | 
(case proc ctxt eta_t' of  | 
| 55031 | 1001  | 
NONE => (cond_tracing' ctxt simp_debug (fn () => "FAILED"); proc_rews ps)  | 
| 15531 | 1002  | 
| SOME raw_thm =>  | 
| 55031 | 1003  | 
(cond_tracing ctxt (fn () =>  | 
| 
55028
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
1004  | 
                    print_thm ctxt ("Procedure " ^ quote name ^ " produced rewrite rule:")
 | 
| 
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
1005  | 
                      ("", raw_thm));
 | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1006  | 
(case rews (mk_procrule ctxt raw_thm) of  | 
| 
55028
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
1007  | 
NONE =>  | 
| 55031 | 1008  | 
(cond_tracing ctxt (fn () =>  | 
| 
55028
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
1009  | 
                        print_term ctxt ("IGNORED result of simproc " ^ quote name ^
 | 
| 
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
1010  | 
" -- does not match") (Thm.term_of t));  | 
| 
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
1011  | 
proc_rews ps)  | 
| 54725 | 1012  | 
| some => some))))  | 
| 10413 | 1013  | 
else proc_rews ps;  | 
| 
38834
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
1014  | 
in  | 
| 
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
1015  | 
(case eta_t of  | 
| 
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
1016  | 
Abs _ $ _ => SOME (Thm.transitive eta_thm (Thm.beta_conversion false eta_t'), skel0)  | 
| 
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
1017  | 
| _ =>  | 
| 
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
1018  | 
(case rews (sort_rrules (Net.match_term rules eta_t)) of  | 
| 
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
1019  | 
NONE => proc_rews (Net.match_term procs eta_t)  | 
| 
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
1020  | 
| some => some))  | 
| 10413 | 1021  | 
end;  | 
1022  | 
||
1023  | 
||
1024  | 
(* conversion to apply a congruence rule to a term *)  | 
|
1025  | 
||
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1026  | 
fun congc prover ctxt maxt cong t =  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1027  | 
let  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1028  | 
val rthm = Thm.incr_indexes (maxt + 1) cong;  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1029  | 
val rlhs = fst (Thm.dest_equals (Drule.strip_imp_concl (cprop_of rthm)));  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1030  | 
val insts = Thm.match (rlhs, t)  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1031  | 
(* Thm.match can raise Pattern.MATCH;  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1032  | 
is handled when congc is called *)  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1033  | 
val thm' = Thm.instantiate insts (Thm.rename_boundvars (term_of rlhs) (term_of t) rthm);  | 
| 
55028
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
1034  | 
val _ =  | 
| 55031 | 1035  | 
      cond_tracing ctxt (fn () => print_thm ctxt "Applying congruence rule:" ("", thm'));
 | 
1036  | 
    fun err (msg, thm) = (cond_tracing ctxt (fn () => print_thm ctxt msg ("", thm)); NONE);
 | 
|
| 
38834
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
1037  | 
in  | 
| 
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
1038  | 
(case prover thm' of  | 
| 
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
1039  | 
      NONE => err ("Congruence proof failed.  Could not prove", thm')
 | 
| 
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
1040  | 
| SOME thm2 =>  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1041  | 
(case check_conv ctxt true (Drule.beta_eta_conversion t) thm2 of  | 
| 15531 | 1042  | 
          NONE => err ("Congruence proof failed.  Should not have proved", thm2)
 | 
1043  | 
| SOME thm2' =>  | 
|
| 
22902
 
ac833b4bb7ee
moved some Drule operations to Thm (see more_thm.ML);
 
wenzelm 
parents: 
22892 
diff
changeset
 | 
1044  | 
if op aconv (pairself term_of (Thm.dest_equals (cprop_of thm2')))  | 
| 
38834
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
1045  | 
then NONE else SOME thm2'))  | 
| 10413 | 1046  | 
end;  | 
1047  | 
||
1048  | 
val (cA, (cB, cC)) =  | 
|
| 
22902
 
ac833b4bb7ee
moved some Drule operations to Thm (see more_thm.ML);
 
wenzelm 
parents: 
22892 
diff
changeset
 | 
1049  | 
apsnd Thm.dest_equals (Thm.dest_implies (hd (cprems_of Drule.imp_cong)));  | 
| 10413 | 1050  | 
|
| 15531 | 1051  | 
fun transitive1 NONE NONE = NONE  | 
1052  | 
| transitive1 (SOME thm1) NONE = SOME thm1  | 
|
1053  | 
| transitive1 NONE (SOME thm2) = SOME thm2  | 
|
| 54725 | 1054  | 
| transitive1 (SOME thm1) (SOME thm2) = SOME (Thm.transitive thm1 thm2);  | 
| 10413 | 1055  | 
|
| 15531 | 1056  | 
fun transitive2 thm = transitive1 (SOME thm);  | 
1057  | 
fun transitive3 thm = transitive1 thm o SOME;  | 
|
| 
13607
 
6908230623a3
Completely reimplemented mutual simplification of premises.
 
berghofe 
parents: 
13569 
diff
changeset
 | 
1058  | 
|
| 52091 | 1059  | 
fun bottomc ((simprem, useprem, mutsimp), prover, maxidx) =  | 
| 10413 | 1060  | 
let  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1061  | 
fun botc skel ctxt t =  | 
| 54725 | 1062  | 
if is_Var skel then NONE  | 
1063  | 
else  | 
|
1064  | 
(case subc skel ctxt t of  | 
|
1065  | 
some as SOME thm1 =>  | 
|
1066  | 
(case rewritec (prover, maxidx) ctxt (Thm.rhs_of thm1) of  | 
|
1067  | 
SOME (thm2, skel2) =>  | 
|
1068  | 
transitive2 (Thm.transitive thm1 thm2)  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1069  | 
(botc skel2 ctxt (Thm.rhs_of thm2))  | 
| 54725 | 1070  | 
| NONE => some)  | 
1071  | 
| NONE =>  | 
|
1072  | 
(case rewritec (prover, maxidx) ctxt t of  | 
|
1073  | 
SOME (thm2, skel2) => transitive2 thm2  | 
|
1074  | 
(botc skel2 ctxt (Thm.rhs_of thm2))  | 
|
1075  | 
| NONE => NONE))  | 
|
| 10413 | 1076  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1077  | 
and try_botc ctxt t =  | 
| 54725 | 1078  | 
(case botc skel0 ctxt t of  | 
1079  | 
SOME trec1 => trec1  | 
|
1080  | 
| NONE => Thm.reflexive t)  | 
|
| 10413 | 1081  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1082  | 
and subc skel ctxt t0 =  | 
| 
55014
 
a93f496f6c30
general notion of auxiliary bounds within context;
 
wenzelm 
parents: 
55000 
diff
changeset
 | 
1083  | 
        let val Simpset (_, {congs, ...}) = simpset_of ctxt in
 | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1084  | 
(case term_of t0 of  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1085  | 
Abs (a, T, _) =>  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1086  | 
let  | 
| 55635 | 1087  | 
val (v, ctxt') = Variable.next_bound (a, T) ctxt;  | 
1088  | 
val b = #1 (Term.dest_Free v);  | 
|
1089  | 
val (v', t') = Thm.dest_abs (SOME b) t0;  | 
|
1090  | 
val b' = #1 (Term.dest_Free (term_of v'));  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1091  | 
val _ =  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1092  | 
if b <> b' then  | 
| 55635 | 1093  | 
                        warning ("Bad Simplifier context: renamed bound variable " ^
 | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1094  | 
quote b ^ " to " ^ quote b' ^ Position.here (Position.thread_data ()))  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1095  | 
else ();  | 
| 54725 | 1096  | 
val skel' = (case skel of Abs (_, _, sk) => sk | _ => skel0);  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1097  | 
in  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1098  | 
(case botc skel' ctxt' t' of  | 
| 55635 | 1099  | 
SOME thm => SOME (Thm.abstract_rule a v' thm)  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1100  | 
| NONE => NONE)  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1101  | 
end  | 
| 54725 | 1102  | 
| t $ _ =>  | 
1103  | 
(case t of  | 
|
| 56245 | 1104  | 
                Const ("Pure.imp", _) $ _  => impc t0 ctxt
 | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1105  | 
| Abs _ =>  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1106  | 
let val thm = Thm.beta_conversion false t0  | 
| 54725 | 1107  | 
in  | 
1108  | 
(case subc skel0 ctxt (Thm.rhs_of thm) of  | 
|
1109  | 
NONE => SOME thm  | 
|
1110  | 
| SOME thm' => SOME (Thm.transitive thm thm'))  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1111  | 
end  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1112  | 
| _ =>  | 
| 54727 | 1113  | 
let  | 
1114  | 
fun appc () =  | 
|
1115  | 
let  | 
|
1116  | 
val (tskel, uskel) =  | 
|
1117  | 
(case skel of  | 
|
1118  | 
tskel $ uskel => (tskel, uskel)  | 
|
1119  | 
| _ => (skel0, skel0));  | 
|
1120  | 
val (ct, cu) = Thm.dest_comb t0;  | 
|
1121  | 
in  | 
|
1122  | 
(case botc tskel ctxt ct of  | 
|
1123  | 
SOME thm1 =>  | 
|
1124  | 
(case botc uskel ctxt cu of  | 
|
1125  | 
SOME thm2 => SOME (Thm.combination thm1 thm2)  | 
|
1126  | 
| NONE => SOME (Thm.combination thm1 (Thm.reflexive cu)))  | 
|
1127  | 
| NONE =>  | 
|
1128  | 
(case botc uskel ctxt cu of  | 
|
1129  | 
SOME thm1 => SOME (Thm.combination (Thm.reflexive ct) thm1)  | 
|
1130  | 
| NONE => NONE))  | 
|
1131  | 
end;  | 
|
1132  | 
val (h, ts) = strip_comb t;  | 
|
| 54725 | 1133  | 
in  | 
1134  | 
(case cong_name h of  | 
|
1135  | 
SOME a =>  | 
|
1136  | 
(case AList.lookup (op =) (fst congs) a of  | 
|
1137  | 
NONE => appc ()  | 
|
1138  | 
| SOME cong =>  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1139  | 
(*post processing: some partial applications h t1 ... tj, j <= length ts,  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1140  | 
may be a redex. Example: map (%x. x) = (%xs. xs) wrt map_cong*)  | 
| 54725 | 1141  | 
(let  | 
1142  | 
val thm = congc (prover ctxt) ctxt maxidx cong t0;  | 
|
1143  | 
val t = the_default t0 (Option.map Thm.rhs_of thm);  | 
|
1144  | 
val (cl, cr) = Thm.dest_comb t  | 
|
1145  | 
                              val dVar = Var(("", 0), dummyT)
 | 
|
1146  | 
val skel =  | 
|
1147  | 
list_comb (h, replicate (length ts) dVar)  | 
|
1148  | 
in  | 
|
1149  | 
(case botc skel ctxt cl of  | 
|
1150  | 
NONE => thm  | 
|
1151  | 
| SOME thm' =>  | 
|
1152  | 
transitive3 thm (Thm.combination thm' (Thm.reflexive cr)))  | 
|
1153  | 
end handle Pattern.MATCH => appc ()))  | 
|
1154  | 
| _ => appc ())  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1155  | 
end)  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1156  | 
| _ => NONE)  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1157  | 
end  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1158  | 
and impc ct ctxt =  | 
| 54725 | 1159  | 
if mutsimp then mut_impc0 [] ct [] [] ctxt  | 
1160  | 
else nonmut_impc ct ctxt  | 
|
| 10413 | 1161  | 
|
| 
54984
 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 
wenzelm 
parents: 
54982 
diff
changeset
 | 
1162  | 
and rules_of_prem prem ctxt =  | 
| 
13607
 
6908230623a3
Completely reimplemented mutual simplification of premises.
 
berghofe 
parents: 
13569 
diff
changeset
 | 
1163  | 
if maxidx_of_term (term_of prem) <> ~1  | 
| 
55028
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
1164  | 
then  | 
| 55031 | 1165  | 
(cond_tracing ctxt (fn () =>  | 
| 
55028
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
1166  | 
print_term ctxt "Cannot add premise as rewrite rule because it contains (type) unknowns:"  | 
| 
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
1167  | 
(term_of prem));  | 
| 
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
1168  | 
(([], NONE), ctxt))  | 
| 
13607
 
6908230623a3
Completely reimplemented mutual simplification of premises.
 
berghofe 
parents: 
13569 
diff
changeset
 | 
1169  | 
else  | 
| 
54984
 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 
wenzelm 
parents: 
54982 
diff
changeset
 | 
1170  | 
let val (asm, ctxt') = Thm.assume_hyps prem ctxt  | 
| 
 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 
wenzelm 
parents: 
54982 
diff
changeset
 | 
1171  | 
in ((extract_safe_rrules ctxt' asm, SOME asm), ctxt') end  | 
| 10413 | 1172  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1173  | 
and add_rrules (rrss, asms) ctxt =  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1174  | 
(fold o fold) insert_rrule rrss ctxt |> add_prems (map_filter I asms)  | 
| 10413 | 1175  | 
|
| 23178 | 1176  | 
and disch r prem eq =  | 
| 
13607
 
6908230623a3
Completely reimplemented mutual simplification of premises.
 
berghofe 
parents: 
13569 
diff
changeset
 | 
1177  | 
let  | 
| 
22902
 
ac833b4bb7ee
moved some Drule operations to Thm (see more_thm.ML);
 
wenzelm 
parents: 
22892 
diff
changeset
 | 
1178  | 
val (lhs, rhs) = Thm.dest_equals (Thm.cprop_of eq);  | 
| 54727 | 1179  | 
val eq' =  | 
1180  | 
Thm.implies_elim  | 
|
1181  | 
(Thm.instantiate ([], [(cA, prem), (cB, lhs), (cC, rhs)]) Drule.imp_cong)  | 
|
1182  | 
(Thm.implies_intr prem eq);  | 
|
| 54725 | 1183  | 
in  | 
1184  | 
if not r then eq'  | 
|
1185  | 
else  | 
|
1186  | 
let  | 
|
1187  | 
val (prem', concl) = Thm.dest_implies lhs;  | 
|
| 54727 | 1188  | 
val (prem'', _) = Thm.dest_implies rhs;  | 
1189  | 
in  | 
|
1190  | 
Thm.transitive  | 
|
1191  | 
(Thm.transitive  | 
|
1192  | 
(Thm.instantiate ([], [(cA, prem'), (cB, prem), (cC, concl)]) Drule.swap_prems_eq)  | 
|
1193  | 
eq')  | 
|
1194  | 
(Thm.instantiate ([], [(cA, prem), (cB, prem''), (cC, concl)]) Drule.swap_prems_eq)  | 
|
| 54725 | 1195  | 
end  | 
| 10413 | 1196  | 
end  | 
1197  | 
||
| 
13607
 
6908230623a3
Completely reimplemented mutual simplification of premises.
 
berghofe 
parents: 
13569 
diff
changeset
 | 
1198  | 
and rebuild [] _ _ _ _ eq = eq  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1199  | 
| rebuild (prem :: prems) concl (_ :: rrss) (_ :: asms) ctxt eq =  | 
| 
13607
 
6908230623a3
Completely reimplemented mutual simplification of premises.
 
berghofe 
parents: 
13569 
diff
changeset
 | 
1200  | 
let  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1201  | 
val ctxt' = add_rrules (rev rrss, rev asms) ctxt;  | 
| 
13607
 
6908230623a3
Completely reimplemented mutual simplification of premises.
 
berghofe 
parents: 
13569 
diff
changeset
 | 
1202  | 
val concl' =  | 
| 
22902
 
ac833b4bb7ee
moved some Drule operations to Thm (see more_thm.ML);
 
wenzelm 
parents: 
22892 
diff
changeset
 | 
1203  | 
Drule.mk_implies (prem, the_default concl (Option.map Thm.rhs_of eq));  | 
| 54727 | 1204  | 
val dprem = Option.map (disch false prem);  | 
| 
38834
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
1205  | 
in  | 
| 52091 | 1206  | 
(case rewritec (prover, maxidx) ctxt' concl' of  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1207  | 
NONE => rebuild prems concl' rrss asms ctxt (dprem eq)  | 
| 54727 | 1208  | 
| SOME (eq', _) =>  | 
1209  | 
transitive2 (fold (disch false) prems (the (transitive3 (dprem eq) eq')))  | 
|
1210  | 
(mut_impc0 (rev prems) (Thm.rhs_of eq') (rev rrss) (rev asms) ctxt))  | 
|
| 
13607
 
6908230623a3
Completely reimplemented mutual simplification of premises.
 
berghofe 
parents: 
13569 
diff
changeset
 | 
1211  | 
end  | 
| 15023 | 1212  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1213  | 
and mut_impc0 prems concl rrss asms ctxt =  | 
| 
13607
 
6908230623a3
Completely reimplemented mutual simplification of premises.
 
berghofe 
parents: 
13569 
diff
changeset
 | 
1214  | 
let  | 
| 
 
6908230623a3
Completely reimplemented mutual simplification of premises.
 
berghofe 
parents: 
13569 
diff
changeset
 | 
1215  | 
val prems' = strip_imp_prems concl;  | 
| 
54984
 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 
wenzelm 
parents: 
54982 
diff
changeset
 | 
1216  | 
val ((rrss', asms'), ctxt') = fold_map rules_of_prem prems' ctxt |>> split_list;  | 
| 
38834
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
1217  | 
in  | 
| 
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
1218  | 
mut_impc (prems @ prems') (strip_imp_concl concl) (rrss @ rrss')  | 
| 
54984
 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 
wenzelm 
parents: 
54982 
diff
changeset
 | 
1219  | 
(asms @ asms') [] [] [] [] ctxt' ~1 ~1  | 
| 
13607
 
6908230623a3
Completely reimplemented mutual simplification of premises.
 
berghofe 
parents: 
13569 
diff
changeset
 | 
1220  | 
end  | 
| 15023 | 1221  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1222  | 
and mut_impc [] concl [] [] prems' rrss' asms' eqns ctxt changed k =  | 
| 33245 | 1223  | 
transitive1 (fold (fn (eq1, prem) => fn eq2 => transitive1 eq1  | 
1224  | 
(Option.map (disch false prem) eq2)) (eqns ~~ prems') NONE)  | 
|
| 
13607
 
6908230623a3
Completely reimplemented mutual simplification of premises.
 
berghofe 
parents: 
13569 
diff
changeset
 | 
1225  | 
(if changed > 0 then  | 
| 
 
6908230623a3
Completely reimplemented mutual simplification of premises.
 
berghofe 
parents: 
13569 
diff
changeset
 | 
1226  | 
mut_impc (rev prems') concl (rev rrss') (rev asms')  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1227  | 
[] [] [] [] ctxt ~1 changed  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1228  | 
else rebuild prems' concl rrss' asms' ctxt  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1229  | 
(botc skel0 (add_rrules (rev rrss', rev asms') ctxt) concl))  | 
| 
13607
 
6908230623a3
Completely reimplemented mutual simplification of premises.
 
berghofe 
parents: 
13569 
diff
changeset
 | 
1230  | 
|
| 
 
6908230623a3
Completely reimplemented mutual simplification of premises.
 
berghofe 
parents: 
13569 
diff
changeset
 | 
1231  | 
| mut_impc (prem :: prems) concl (rrs :: rrss) (asm :: asms)  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1232  | 
prems' rrss' asms' eqns ctxt changed k =  | 
| 54725 | 1233  | 
(case (if k = 0 then NONE else botc skel0 (add_rrules  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1234  | 
(rev rrss' @ rrss, rev asms' @ asms) ctxt) prem) of  | 
| 15531 | 1235  | 
NONE => mut_impc prems concl rrss asms (prem :: prems')  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1236  | 
(rrs :: rrss') (asm :: asms') (NONE :: eqns) ctxt changed  | 
| 
13607
 
6908230623a3
Completely reimplemented mutual simplification of premises.
 
berghofe 
parents: 
13569 
diff
changeset
 | 
1237  | 
(if k = 0 then 0 else k - 1)  | 
| 54725 | 1238  | 
| SOME eqn =>  | 
| 
13607
 
6908230623a3
Completely reimplemented mutual simplification of premises.
 
berghofe 
parents: 
13569 
diff
changeset
 | 
1239  | 
let  | 
| 
22902
 
ac833b4bb7ee
moved some Drule operations to Thm (see more_thm.ML);
 
wenzelm 
parents: 
22892 
diff
changeset
 | 
1240  | 
val prem' = Thm.rhs_of eqn;  | 
| 
13607
 
6908230623a3
Completely reimplemented mutual simplification of premises.
 
berghofe 
parents: 
13569 
diff
changeset
 | 
1241  | 
val tprems = map term_of prems;  | 
| 33029 | 1242  | 
val i = 1 + fold Integer.max (map (fn p =>  | 
| 44058 | 1243  | 
find_index (fn q => q aconv p) tprems) (Thm.hyps_of eqn)) ~1;  | 
| 
54984
 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 
wenzelm 
parents: 
54982 
diff
changeset
 | 
1244  | 
val ((rrs', asm'), ctxt') = rules_of_prem prem' ctxt;  | 
| 54725 | 1245  | 
in  | 
1246  | 
mut_impc prems concl rrss asms (prem' :: prems')  | 
|
1247  | 
(rrs' :: rrss') (asm' :: asms')  | 
|
1248  | 
(SOME (fold_rev (disch true)  | 
|
1249  | 
(take i prems)  | 
|
1250  | 
(Drule.imp_cong_rule eqn (Thm.reflexive (Drule.list_implies  | 
|
1251  | 
(drop i prems, concl))))) :: eqns)  | 
|
| 
54984
 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 
wenzelm 
parents: 
54982 
diff
changeset
 | 
1252  | 
ctxt' (length prems') ~1  | 
| 54725 | 1253  | 
end)  | 
| 
13607
 
6908230623a3
Completely reimplemented mutual simplification of premises.
 
berghofe 
parents: 
13569 
diff
changeset
 | 
1254  | 
|
| 54725 | 1255  | 
(*legacy code -- only for backwards compatibility*)  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1256  | 
and nonmut_impc ct ctxt =  | 
| 
38834
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
1257  | 
let  | 
| 
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
1258  | 
val (prem, conc) = Thm.dest_implies ct;  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1259  | 
val thm1 = if simprem then botc skel0 ctxt prem else NONE;  | 
| 
38834
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
1260  | 
val prem1 = the_default prem (Option.map Thm.rhs_of thm1);  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1261  | 
val ctxt1 =  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1262  | 
if not useprem then ctxt  | 
| 
54984
 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 
wenzelm 
parents: 
54982 
diff
changeset
 | 
1263  | 
else  | 
| 
 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 
wenzelm 
parents: 
54982 
diff
changeset
 | 
1264  | 
let val ((rrs, asm), ctxt') = rules_of_prem prem1 ctxt  | 
| 
 
da70ab8531f4
more elementary management of declared hyps, below structure Assumption;
 
wenzelm 
parents: 
54982 
diff
changeset
 | 
1265  | 
in add_rrules ([rrs], [asm]) ctxt' end;  | 
| 
38834
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
1266  | 
in  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1267  | 
(case botc skel0 ctxt1 conc of  | 
| 
38834
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
1268  | 
NONE =>  | 
| 
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
1269  | 
(case thm1 of  | 
| 
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
1270  | 
NONE => NONE  | 
| 
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
1271  | 
| SOME thm1' => SOME (Drule.imp_cong_rule thm1' (Thm.reflexive conc)))  | 
| 
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
1272  | 
| SOME thm2 =>  | 
| 
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
1273  | 
let val thm2' = disch false prem1 thm2 in  | 
| 
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
1274  | 
(case thm1 of  | 
| 
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
1275  | 
NONE => SOME thm2'  | 
| 
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
1276  | 
| SOME thm1' =>  | 
| 36944 | 1277  | 
SOME (Thm.transitive (Drule.imp_cong_rule thm1' (Thm.reflexive conc)) thm2'))  | 
| 
38834
 
658fcba35ed7
more careful treatment of context visibility flag wrt. spurious warnings;
 
wenzelm 
parents: 
38715 
diff
changeset
 | 
1278  | 
end)  | 
| 54725 | 1279  | 
end;  | 
| 10413 | 1280  | 
|
| 54725 | 1281  | 
in try_botc end;  | 
| 10413 | 1282  | 
|
1283  | 
||
| 15023 | 1284  | 
(* Meta-rewriting: rewrites t to u and returns the theorem t==u *)  | 
| 10413 | 1285  | 
|
1286  | 
(*  | 
|
1287  | 
Parameters:  | 
|
1288  | 
mode = (simplify A,  | 
|
1289  | 
use A in simplifying B,  | 
|
1290  | 
use prems of B (if B is again a meta-impl.) to simplify A)  | 
|
1291  | 
when simplifying A ==> B  | 
|
1292  | 
prover: how to solve premises in conditional rewrites and congruences  | 
|
1293  | 
*)  | 
|
1294  | 
||
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1295  | 
fun rewrite_cterm mode prover raw_ctxt raw_ct =  | 
| 
17882
 
b6e44fc46cf0
added set/addloop' for simpset dependent loopers;
 
wenzelm 
parents: 
17756 
diff
changeset
 | 
1296  | 
let  | 
| 
54729
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
1297  | 
val thy = Proof_Context.theory_of raw_ctxt;  | 
| 52091 | 1298  | 
|
| 20260 | 1299  | 
val ct = Thm.adjust_maxidx_cterm ~1 raw_ct;  | 
| 32797 | 1300  | 
    val {maxidx, ...} = Thm.rep_cterm ct;
 | 
| 52091 | 1301  | 
val _ =  | 
1302  | 
Theory.subthy (theory_of_cterm ct, thy) orelse  | 
|
1303  | 
        raise CTERM ("rewrite_cterm: bad background theory", [ct]);
 | 
|
1304  | 
||
| 
54729
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
1305  | 
val ctxt =  | 
| 
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
1306  | 
raw_ctxt  | 
| 
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
1307  | 
|> Context_Position.set_visible false  | 
| 
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
1308  | 
|> inc_simp_depth  | 
| 
55028
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
1309  | 
      |> (fn ctxt => trace_invoke {depth = simp_depth ctxt, term = term_of ct} ctxt);
 | 
| 
54729
 
c5cd7a58cf2d
generic trace operations for main steps of Simplifier;
 
wenzelm 
parents: 
54728 
diff
changeset
 | 
1310  | 
|
| 
55028
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
1311  | 
val _ =  | 
| 55031 | 1312  | 
cond_tracing ctxt (fn () =>  | 
| 
55028
 
00e849f5b397
clarified Simplifier diagnostics -- simplified ML;
 
wenzelm 
parents: 
55014 
diff
changeset
 | 
1313  | 
print_term ctxt "SIMPLIFIER INVOKED ON THE FOLLOWING TERM:" (term_of ct));  | 
| 52091 | 1314  | 
in bottomc (mode, Option.map Drule.flexflex_unique oo prover, maxidx) ctxt ct end;  | 
| 10413 | 1315  | 
|
| 21708 | 1316  | 
val simple_prover =  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1317  | 
SINGLE o (fn ctxt => ALLGOALS (resolve_tac (prems_of ctxt)));  | 
| 21708 | 1318  | 
|
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
54731 
diff
changeset
 | 
1319  | 
fun rewrite _ _ [] = Thm.reflexive  | 
| 
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
54731 
diff
changeset
 | 
1320  | 
| rewrite ctxt full thms =  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1321  | 
rewrite_cterm (full, false, false) simple_prover  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
54731 
diff
changeset
 | 
1322  | 
(empty_simpset ctxt addsimps thms);  | 
| 11672 | 1323  | 
|
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
54731 
diff
changeset
 | 
1324  | 
fun rewrite_rule ctxt = Conv.fconv_rule o rewrite ctxt true;  | 
| 21708 | 1325  | 
|
| 15023 | 1326  | 
(*simple term rewriting -- no proof*)  | 
| 16458 | 1327  | 
fun rewrite_term thy rules procs =  | 
| 17203 | 1328  | 
Pattern.rewrite_term thy (map decomp_simp' rules) procs;  | 
| 15023 | 1329  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1330  | 
fun rewrite_thm mode prover ctxt = Conv.fconv_rule (rewrite_cterm mode prover ctxt);  | 
| 10413 | 1331  | 
|
| 
23536
 
60a1672e298e
moved (asm_)rewrite_goal_tac from goal.ML to meta_simplifier.ML (no longer depends on SELECT_GOAL);
 
wenzelm 
parents: 
23221 
diff
changeset
 | 
1332  | 
(*Rewrite the subgoals of a proof state (represented by a theorem)*)  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
54731 
diff
changeset
 | 
1333  | 
fun rewrite_goals_rule ctxt thms th =  | 
| 23584 | 1334  | 
Conv.fconv_rule (Conv.prems_conv ~1 (rewrite_cterm (true, true, true) simple_prover  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
54731 
diff
changeset
 | 
1335  | 
(empty_simpset ctxt addsimps thms))) th;  | 
| 10413 | 1336  | 
|
| 
20228
 
e0f9e8a6556b
moved Goal.norm_hhf(_protect) to meta_simplifier.ML (pervasive);
 
wenzelm 
parents: 
20197 
diff
changeset
 | 
1337  | 
|
| 21708 | 1338  | 
(** meta-rewriting tactics **)  | 
1339  | 
||
| 
28839
 
32d498cf7595
eliminated rewrite_tac/fold_tac, which are not well-formed tactics due to change of main conclusion;
 
wenzelm 
parents: 
28620 
diff
changeset
 | 
1340  | 
(*Rewrite all subgoals*)  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
54731 
diff
changeset
 | 
1341  | 
fun rewrite_goals_tac ctxt defs = PRIMITIVE (rewrite_goals_rule ctxt defs);  | 
| 21708 | 1342  | 
|
| 
28839
 
32d498cf7595
eliminated rewrite_tac/fold_tac, which are not well-formed tactics due to change of main conclusion;
 
wenzelm 
parents: 
28620 
diff
changeset
 | 
1343  | 
(*Rewrite one subgoal*)  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1344  | 
fun generic_rewrite_goal_tac mode prover_tac ctxt i thm =  | 
| 
25203
 
e5b2dd8db7c8
asm_rewrite_goal_tac: avoiding PRIMITIVE lets informative exceptions (from simprocs) get through;
 
wenzelm 
parents: 
24707 
diff
changeset
 | 
1345  | 
if 0 < i andalso i <= Thm.nprems_of thm then  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1346  | 
Seq.single (Conv.gconv_rule (rewrite_cterm mode (SINGLE o prover_tac) ctxt) i thm)  | 
| 
25203
 
e5b2dd8db7c8
asm_rewrite_goal_tac: avoiding PRIMITIVE lets informative exceptions (from simprocs) get through;
 
wenzelm 
parents: 
24707 
diff
changeset
 | 
1347  | 
else Seq.empty;  | 
| 
23536
 
60a1672e298e
moved (asm_)rewrite_goal_tac from goal.ML to meta_simplifier.ML (no longer depends on SELECT_GOAL);
 
wenzelm 
parents: 
23221 
diff
changeset
 | 
1348  | 
|
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
54731 
diff
changeset
 | 
1349  | 
fun rewrite_goal_tac ctxt rews =  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1350  | 
generic_rewrite_goal_tac (true, false, false) (K no_tac)  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
54731 
diff
changeset
 | 
1351  | 
(empty_simpset ctxt addsimps rews);  | 
| 
23536
 
60a1672e298e
moved (asm_)rewrite_goal_tac from goal.ML to meta_simplifier.ML (no longer depends on SELECT_GOAL);
 
wenzelm 
parents: 
23221 
diff
changeset
 | 
1352  | 
|
| 46707 | 1353  | 
(*Prunes all redundant parameters from the proof state by rewriting.*)  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
54731 
diff
changeset
 | 
1354  | 
fun prune_params_tac ctxt = rewrite_goals_tac ctxt [Drule.triv_forall_equality];  | 
| 21708 | 1355  | 
|
1356  | 
||
1357  | 
(* for folding definitions, handling critical pairs *)  | 
|
1358  | 
||
1359  | 
(*The depth of nesting in a term*)  | 
|
| 32797 | 1360  | 
fun term_depth (Abs (_, _, t)) = 1 + term_depth t  | 
1361  | 
| term_depth (f $ t) = 1 + Int.max (term_depth f, term_depth t)  | 
|
| 21708 | 1362  | 
| term_depth _ = 0;  | 
1363  | 
||
1364  | 
val lhs_of_thm = #1 o Logic.dest_equals o prop_of;  | 
|
1365  | 
||
1366  | 
(*folding should handle critical pairs! E.g. K == Inl(0), S == Inr(Inl(0))  | 
|
1367  | 
Returns longest lhs first to avoid folding its subexpressions.*)  | 
|
1368  | 
fun sort_lhs_depths defs =  | 
|
1369  | 
let val keylist = AList.make (term_depth o lhs_of_thm) defs  | 
|
1370  | 
val keys = sort_distinct (rev_order o int_ord) (map #2 keylist)  | 
|
1371  | 
in map (AList.find (op =) keylist) keys end;  | 
|
1372  | 
||
| 36944 | 1373  | 
val rev_defs = sort_lhs_depths o map Thm.symmetric;  | 
| 21708 | 1374  | 
|
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
54731 
diff
changeset
 | 
1375  | 
fun fold_rule ctxt defs = fold (rewrite_rule ctxt) (rev_defs defs);  | 
| 
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
54731 
diff
changeset
 | 
1376  | 
fun fold_goals_tac ctxt defs = EVERY (map (rewrite_goals_tac ctxt) (rev_defs defs));  | 
| 21708 | 1377  | 
|
1378  | 
||
| 
20228
 
e0f9e8a6556b
moved Goal.norm_hhf(_protect) to meta_simplifier.ML (pervasive);
 
wenzelm 
parents: 
20197 
diff
changeset
 | 
1379  | 
(* HHF normal form: !! before ==>, outermost !! generalized *)  | 
| 
 
e0f9e8a6556b
moved Goal.norm_hhf(_protect) to meta_simplifier.ML (pervasive);
 
wenzelm 
parents: 
20197 
diff
changeset
 | 
1380  | 
|
| 
 
e0f9e8a6556b
moved Goal.norm_hhf(_protect) to meta_simplifier.ML (pervasive);
 
wenzelm 
parents: 
20197 
diff
changeset
 | 
1381  | 
local  | 
| 
 
e0f9e8a6556b
moved Goal.norm_hhf(_protect) to meta_simplifier.ML (pervasive);
 
wenzelm 
parents: 
20197 
diff
changeset
 | 
1382  | 
|
| 
54883
 
dd04a8b654fc
proper context for norm_hhf and derived operations;
 
wenzelm 
parents: 
54742 
diff
changeset
 | 
1383  | 
fun gen_norm_hhf ss ctxt th =  | 
| 21565 | 1384  | 
(if Drule.is_norm_hhf (Thm.prop_of th) then th  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1385  | 
else  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1386  | 
Conv.fconv_rule  | 
| 
54883
 
dd04a8b654fc
proper context for norm_hhf and derived operations;
 
wenzelm 
parents: 
54742 
diff
changeset
 | 
1387  | 
(rewrite_cterm (true, false, false) (K (K NONE)) (put_simpset ss ctxt)) th)  | 
| 21565 | 1388  | 
|> Thm.adjust_maxidx_thm ~1  | 
1389  | 
|> Drule.gen_all;  | 
|
| 
20228
 
e0f9e8a6556b
moved Goal.norm_hhf(_protect) to meta_simplifier.ML (pervasive);
 
wenzelm 
parents: 
20197 
diff
changeset
 | 
1390  | 
|
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1391  | 
val hhf_ss =  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1392  | 
simpset_of (empty_simpset (Context.proof_of (Context.the_thread_data ()))  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1393  | 
addsimps Drule.norm_hhf_eqs);  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1394  | 
|
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1395  | 
val hhf_protect_ss =  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1396  | 
simpset_of (empty_simpset (Context.proof_of (Context.the_thread_data ()))  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1397  | 
addsimps Drule.norm_hhf_eqs |> add_eqcong Drule.protect_cong);  | 
| 
20228
 
e0f9e8a6556b
moved Goal.norm_hhf(_protect) to meta_simplifier.ML (pervasive);
 
wenzelm 
parents: 
20197 
diff
changeset
 | 
1398  | 
|
| 
 
e0f9e8a6556b
moved Goal.norm_hhf(_protect) to meta_simplifier.ML (pervasive);
 
wenzelm 
parents: 
20197 
diff
changeset
 | 
1399  | 
in  | 
| 
 
e0f9e8a6556b
moved Goal.norm_hhf(_protect) to meta_simplifier.ML (pervasive);
 
wenzelm 
parents: 
20197 
diff
changeset
 | 
1400  | 
|
| 26424 | 1401  | 
val norm_hhf = gen_norm_hhf hhf_ss;  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51611 
diff
changeset
 | 
1402  | 
val norm_hhf_protect = gen_norm_hhf hhf_protect_ss;  | 
| 
20228
 
e0f9e8a6556b
moved Goal.norm_hhf(_protect) to meta_simplifier.ML (pervasive);
 
wenzelm 
parents: 
20197 
diff
changeset
 | 
1403  | 
|
| 
 
e0f9e8a6556b
moved Goal.norm_hhf(_protect) to meta_simplifier.ML (pervasive);
 
wenzelm 
parents: 
20197 
diff
changeset
 | 
1404  | 
end;  | 
| 
 
e0f9e8a6556b
moved Goal.norm_hhf(_protect) to meta_simplifier.ML (pervasive);
 
wenzelm 
parents: 
20197 
diff
changeset
 | 
1405  | 
|
| 10413 | 1406  | 
end;  | 
1407  | 
||
| 
41228
 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 
wenzelm 
parents: 
41227 
diff
changeset
 | 
1408  | 
structure Basic_Meta_Simplifier: BASIC_RAW_SIMPLIFIER = Raw_Simplifier;  | 
| 32738 | 1409  | 
open Basic_Meta_Simplifier;  |