| author | wenzelm | 
| Tue, 12 Aug 2025 11:19:08 +0200 | |
| changeset 82996 | 4a77ce6d4e07 | 
| parent 82887 | e5db7672d192 | 
| permissions | -rw-r--r-- | 
| 9938 | 1  | 
(* Title: Provers/classical.ML  | 
2  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
|
| 82852 | 3  | 
Author: Makarius  | 
| 0 | 4  | 
|
5  | 
Theorem prover for classical reasoning, including predicate calculus, set  | 
|
6  | 
theory, etc.  | 
|
7  | 
||
| 60943 | 8  | 
Rules must be classified as intro, elim, safe, unsafe.  | 
| 0 | 9  | 
|
10  | 
A rule is unsafe unless it can be applied blindly without harmful results.  | 
|
11  | 
For a rule to be safe, its premises and conclusion should be logically  | 
|
12  | 
equivalent. There should be no variables in the premises that are not in  | 
|
13  | 
the conclusion.  | 
|
| 82852 | 14  | 
|
15  | 
Rules are maintained in "canonical reverse order", meaning that later  | 
|
16  | 
declarations take precedence.  | 
|
| 0 | 17  | 
*)  | 
18  | 
||
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
19  | 
(*higher precedence than := facilitates use of references*)  | 
| 
12376
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
20  | 
infix 4 addSIs addSEs addSDs addIs addEs addDs delrules  | 
| 4651 | 21  | 
addSWrapper delSWrapper addWrapper delWrapper  | 
| 
11181
 
d04f57b91166
renamed addaltern to addafter, addSaltern to addSafter
 
oheimb 
parents: 
11168 
diff
changeset
 | 
22  | 
addSbefore addSafter addbefore addafter  | 
| 5523 | 23  | 
addD2 addE2 addSD2 addSE2;  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
24  | 
|
| 0 | 25  | 
signature CLASSICAL_DATA =  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
26  | 
sig  | 
| 82852 | 27  | 
val imp_elim: thm (* P \<longrightarrow> Q \<Longrightarrow> (\<not> R \<Longrightarrow> P) \<Longrightarrow> (Q \<Longrightarrow> R) \<Longrightarrow> R *)  | 
28  | 
val not_elim: thm (* \<not> P \<Longrightarrow> P \<Longrightarrow> R *)  | 
|
29  | 
val swap: thm (* \<not> P \<Longrightarrow> (\<not> R \<Longrightarrow> P) \<Longrightarrow> R *)  | 
|
30  | 
val classical: thm (* (\<not> P \<Longrightarrow> P) \<Longrightarrow> P *)  | 
|
31  | 
val sizef: thm -> int (*size function for BEST_FIRST, typically size_of_thm*)  | 
|
32  | 
val hyp_subst_tacs: (Proof.context -> int -> tactic) list  | 
|
33  | 
(*optional tactics for substitution in the hypotheses; assumed to be safe!*)  | 
|
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
34  | 
end;  | 
| 0 | 35  | 
|
| 5841 | 36  | 
signature BASIC_CLASSICAL =  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
37  | 
sig  | 
| 42812 | 38  | 
type wrapper = (int -> tactic) -> int -> tactic  | 
| 0 | 39  | 
type claset  | 
| 42793 | 40  | 
val addDs: Proof.context * thm list -> Proof.context  | 
41  | 
val addEs: Proof.context * thm list -> Proof.context  | 
|
42  | 
val addIs: Proof.context * thm list -> Proof.context  | 
|
43  | 
val addSDs: Proof.context * thm list -> Proof.context  | 
|
44  | 
val addSEs: Proof.context * thm list -> Proof.context  | 
|
45  | 
val addSIs: Proof.context * thm list -> Proof.context  | 
|
46  | 
val delrules: Proof.context * thm list -> Proof.context  | 
|
| 
51703
 
f2e92fc0c8aa
modifiers for classical wrappers operate on Proof.context instead of claset;
 
wenzelm 
parents: 
51688 
diff
changeset
 | 
47  | 
val addSWrapper: Proof.context * (string * (Proof.context -> wrapper)) -> Proof.context  | 
| 
 
f2e92fc0c8aa
modifiers for classical wrappers operate on Proof.context instead of claset;
 
wenzelm 
parents: 
51688 
diff
changeset
 | 
48  | 
val delSWrapper: Proof.context * string -> Proof.context  | 
| 
 
f2e92fc0c8aa
modifiers for classical wrappers operate on Proof.context instead of claset;
 
wenzelm 
parents: 
51688 
diff
changeset
 | 
49  | 
val addWrapper: Proof.context * (string * (Proof.context -> wrapper)) -> Proof.context  | 
| 
 
f2e92fc0c8aa
modifiers for classical wrappers operate on Proof.context instead of claset;
 
wenzelm 
parents: 
51688 
diff
changeset
 | 
50  | 
val delWrapper: Proof.context * string -> Proof.context  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51703 
diff
changeset
 | 
51  | 
val addSbefore: Proof.context * (string * (Proof.context -> int -> tactic)) -> Proof.context  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51703 
diff
changeset
 | 
52  | 
val addSafter: Proof.context * (string * (Proof.context -> int -> tactic)) -> Proof.context  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51703 
diff
changeset
 | 
53  | 
val addbefore: Proof.context * (string * (Proof.context -> int -> tactic)) -> Proof.context  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51703 
diff
changeset
 | 
54  | 
val addafter: Proof.context * (string * (Proof.context -> int -> tactic)) -> Proof.context  | 
| 
51703
 
f2e92fc0c8aa
modifiers for classical wrappers operate on Proof.context instead of claset;
 
wenzelm 
parents: 
51688 
diff
changeset
 | 
55  | 
val addD2: Proof.context * (string * thm) -> Proof.context  | 
| 
 
f2e92fc0c8aa
modifiers for classical wrappers operate on Proof.context instead of claset;
 
wenzelm 
parents: 
51688 
diff
changeset
 | 
56  | 
val addE2: Proof.context * (string * thm) -> Proof.context  | 
| 
 
f2e92fc0c8aa
modifiers for classical wrappers operate on Proof.context instead of claset;
 
wenzelm 
parents: 
51688 
diff
changeset
 | 
57  | 
val addSD2: Proof.context * (string * thm) -> Proof.context  | 
| 
 
f2e92fc0c8aa
modifiers for classical wrappers operate on Proof.context instead of claset;
 
wenzelm 
parents: 
51688 
diff
changeset
 | 
58  | 
val addSE2: Proof.context * (string * thm) -> Proof.context  | 
| 42793 | 59  | 
val appSWrappers: Proof.context -> wrapper  | 
60  | 
val appWrappers: Proof.context -> wrapper  | 
|
| 
982
 
4fe0b642b7d5
Addition of wrappers for integration with the simplifier.
 
lcp 
parents: 
747 
diff
changeset
 | 
61  | 
|
| 42790 | 62  | 
val claset_of: Proof.context -> claset  | 
| 42793 | 63  | 
val put_claset: claset -> Proof.context -> Proof.context  | 
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
64  | 
|
| 
51703
 
f2e92fc0c8aa
modifiers for classical wrappers operate on Proof.context instead of claset;
 
wenzelm 
parents: 
51688 
diff
changeset
 | 
65  | 
val map_theory_claset: (Proof.context -> Proof.context) -> theory -> theory  | 
| 
 
f2e92fc0c8aa
modifiers for classical wrappers operate on Proof.context instead of claset;
 
wenzelm 
parents: 
51688 
diff
changeset
 | 
66  | 
|
| 42793 | 67  | 
val fast_tac: Proof.context -> int -> tactic  | 
68  | 
val slow_tac: Proof.context -> int -> tactic  | 
|
69  | 
val astar_tac: Proof.context -> int -> tactic  | 
|
70  | 
val slow_astar_tac: Proof.context -> int -> tactic  | 
|
71  | 
val best_tac: Proof.context -> int -> tactic  | 
|
72  | 
val first_best_tac: Proof.context -> int -> tactic  | 
|
73  | 
val slow_best_tac: Proof.context -> int -> tactic  | 
|
74  | 
val depth_tac: Proof.context -> int -> int -> tactic  | 
|
75  | 
val deepen_tac: Proof.context -> int -> int -> tactic  | 
|
| 
1587
 
e7d8a4957bac
Now provides astar versions (thanks to Norbert Voelker)
 
paulson 
parents: 
1524 
diff
changeset
 | 
76  | 
|
| 
58963
 
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
 
wenzelm 
parents: 
58958 
diff
changeset
 | 
77  | 
val contr_tac: Proof.context -> int -> tactic  | 
| 59970 | 78  | 
val dup_elim: Proof.context -> thm -> thm  | 
| 42790 | 79  | 
val dup_intr: thm -> thm  | 
| 42793 | 80  | 
val dup_step_tac: Proof.context -> int -> tactic  | 
| 58957 | 81  | 
val eq_mp_tac: Proof.context -> int -> tactic  | 
| 60943 | 82  | 
val unsafe_step_tac: Proof.context -> int -> tactic  | 
| 
58963
 
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
 
wenzelm 
parents: 
58958 
diff
changeset
 | 
83  | 
val mp_tac: Proof.context -> int -> tactic  | 
| 42793 | 84  | 
val safe_tac: Proof.context -> tactic  | 
85  | 
val safe_steps_tac: Proof.context -> int -> tactic  | 
|
86  | 
val safe_step_tac: Proof.context -> int -> tactic  | 
|
87  | 
val clarify_tac: Proof.context -> int -> tactic  | 
|
88  | 
val clarify_step_tac: Proof.context -> int -> tactic  | 
|
89  | 
val step_tac: Proof.context -> int -> tactic  | 
|
90  | 
val slow_step_tac: Proof.context -> int -> tactic  | 
|
| 
58963
 
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
 
wenzelm 
parents: 
58958 
diff
changeset
 | 
91  | 
val swap_res_tac: Proof.context -> thm list -> int -> tactic  | 
| 42793 | 92  | 
val inst_step_tac: Proof.context -> int -> tactic  | 
93  | 
val inst0_step_tac: Proof.context -> int -> tactic  | 
|
94  | 
val instp_step_tac: Proof.context -> int -> tactic  | 
|
| 82853 | 95  | 
|
96  | 
val print_claset: Proof.context -> unit  | 
|
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
97  | 
end;  | 
| 1724 | 98  | 
|
| 5841 | 99  | 
signature CLASSICAL =  | 
100  | 
sig  | 
|
101  | 
include BASIC_CLASSICAL  | 
|
| 59970 | 102  | 
val classical_rule: Proof.context -> thm -> thm  | 
| 
82833
 
13a8c49a48a0
clarified signature: more explicit types, notably (thm option) instead of (thm list);
 
wenzelm 
parents: 
82832 
diff
changeset
 | 
103  | 
type rl = thm * thm option  | 
| 
82851
 
7f9c4466c6a5
proper "plain" rule for extra_netpair (amending 8aa1c98b948b): need to avoid flat_rule / Object_Logic.atomize_prems for the sake of "standard" Isar proof, e.g. (line 34 of "$AFP/AWN/Qmsg_Lifting.thy);
 
wenzelm 
parents: 
82849 
diff
changeset
 | 
104  | 
  type info = {rl: rl, dup_rl: rl, plain: thm}
 | 
| 
82842
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
105  | 
type decl = info Bires.decl  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
106  | 
type decls = info Bires.decls  | 
| 
82846
 
c906b8df7bc3
clarified signature: do not expose internal data structures;
 
wenzelm 
parents: 
82845 
diff
changeset
 | 
107  | 
val safe0_netpair_of: Proof.context -> Bires.netpair  | 
| 
 
c906b8df7bc3
clarified signature: do not expose internal data structures;
 
wenzelm 
parents: 
82845 
diff
changeset
 | 
108  | 
val safep_netpair_of: Proof.context -> Bires.netpair  | 
| 
 
c906b8df7bc3
clarified signature: do not expose internal data structures;
 
wenzelm 
parents: 
82845 
diff
changeset
 | 
109  | 
val unsafe_netpair_of: Proof.context -> Bires.netpair  | 
| 
 
c906b8df7bc3
clarified signature: do not expose internal data structures;
 
wenzelm 
parents: 
82845 
diff
changeset
 | 
110  | 
val dup_netpair_of: Proof.context -> Bires.netpair  | 
| 
 
c906b8df7bc3
clarified signature: do not expose internal data structures;
 
wenzelm 
parents: 
82845 
diff
changeset
 | 
111  | 
val extra_netpair_of: Proof.context -> Bires.netpair  | 
| 82845 | 112  | 
val dest_decls: Proof.context -> ((thm * decl) -> bool) -> (thm * decl) list  | 
| 24021 | 113  | 
val get_cs: Context.generic -> claset  | 
114  | 
val map_cs: (claset -> claset) -> Context.generic -> Context.generic  | 
|
| 18728 | 115  | 
val safe_dest: int option -> attribute  | 
116  | 
val safe_elim: int option -> attribute  | 
|
117  | 
val safe_intro: int option -> attribute  | 
|
| 60943 | 118  | 
val unsafe_dest: int option -> attribute  | 
119  | 
val unsafe_elim: int option -> attribute  | 
|
120  | 
val unsafe_intro: int option -> attribute  | 
|
| 18728 | 121  | 
val rule_del: attribute  | 
| 61327 | 122  | 
val rule_tac: Proof.context -> thm list -> thm list -> int -> tactic  | 
123  | 
val standard_tac: Proof.context -> thm list -> tactic  | 
|
| 30513 | 124  | 
val cla_modifiers: Method.modifier parser list  | 
| 42793 | 125  | 
val cla_method:  | 
126  | 
(Proof.context -> tactic) -> (Proof.context -> Proof.method) context_parser  | 
|
127  | 
val cla_method':  | 
|
128  | 
(Proof.context -> int -> tactic) -> (Proof.context -> Proof.method) context_parser  | 
|
| 5841 | 129  | 
end;  | 
130  | 
||
| 0 | 131  | 
|
| 42799 | 132  | 
functor Classical(Data: CLASSICAL_DATA): CLASSICAL =  | 
| 0 | 133  | 
struct  | 
134  | 
||
| 82852 | 135  | 
(** Support for classical reasoning **)  | 
| 
18534
 
6799b38ed872
added classical_rule, which replaces Data.make_elim;
 
wenzelm 
parents: 
18374 
diff
changeset
 | 
136  | 
|
| 82852 | 137  | 
(* classical elimination rules *)  | 
138  | 
||
139  | 
(*Classical reasoning requires stronger elimination rules. For  | 
|
140  | 
instance, make_elim of Pure transforms the HOL rule injD into  | 
|
| 
18534
 
6799b38ed872
added classical_rule, which replaces Data.make_elim;
 
wenzelm 
parents: 
18374 
diff
changeset
 | 
141  | 
|
| 82852 | 142  | 
\<lbrakk>inj f; f x = f y; x = y \<Longrightarrow> PROP W\<rbrakk> \<Longrightarrow> PROP W  | 
| 
18534
 
6799b38ed872
added classical_rule, which replaces Data.make_elim;
 
wenzelm 
parents: 
18374 
diff
changeset
 | 
143  | 
|
| 82852 | 144  | 
Such rules can cause fast_tac to fail and blast_tac to report "PROOF  | 
145  | 
FAILED"; classical_rule will strengthen this to  | 
|
| 
18534
 
6799b38ed872
added classical_rule, which replaces Data.make_elim;
 
wenzelm 
parents: 
18374 
diff
changeset
 | 
146  | 
|
| 82852 | 147  | 
\<lbrakk>inj f; \<not> W \<Longrightarrow> f x = f y; x = y \<Longrightarrow> W\<rbrakk> \<Longrightarrow> W  | 
| 
18534
 
6799b38ed872
added classical_rule, which replaces Data.make_elim;
 
wenzelm 
parents: 
18374 
diff
changeset
 | 
148  | 
*)  | 
| 
 
6799b38ed872
added classical_rule, which replaces Data.make_elim;
 
wenzelm 
parents: 
18374 
diff
changeset
 | 
149  | 
|
| 59970 | 150  | 
fun classical_rule ctxt rule =  | 
151  | 
if is_some (Object_Logic.elim_concl ctxt rule) then  | 
|
| 
18534
 
6799b38ed872
added classical_rule, which replaces Data.make_elim;
 
wenzelm 
parents: 
18374 
diff
changeset
 | 
152  | 
let  | 
| 60817 | 153  | 
val thy = Proof_Context.theory_of ctxt;  | 
| 42792 | 154  | 
val rule' = rule RS Data.classical;  | 
| 
18534
 
6799b38ed872
added classical_rule, which replaces Data.make_elim;
 
wenzelm 
parents: 
18374 
diff
changeset
 | 
155  | 
val concl' = Thm.concl_of rule';  | 
| 
 
6799b38ed872
added classical_rule, which replaces Data.make_elim;
 
wenzelm 
parents: 
18374 
diff
changeset
 | 
156  | 
fun redundant_hyp goal =  | 
| 19257 | 157  | 
concl' aconv Logic.strip_assums_concl goal orelse  | 
| 
18534
 
6799b38ed872
added classical_rule, which replaces Data.make_elim;
 
wenzelm 
parents: 
18374 
diff
changeset
 | 
158  | 
(case Logic.strip_assums_hyp goal of  | 
| 
 
6799b38ed872
added classical_rule, which replaces Data.make_elim;
 
wenzelm 
parents: 
18374 
diff
changeset
 | 
159  | 
hyp :: hyps => exists (fn t => t aconv hyp) hyps  | 
| 
 
6799b38ed872
added classical_rule, which replaces Data.make_elim;
 
wenzelm 
parents: 
18374 
diff
changeset
 | 
160  | 
| _ => false);  | 
| 
 
6799b38ed872
added classical_rule, which replaces Data.make_elim;
 
wenzelm 
parents: 
18374 
diff
changeset
 | 
161  | 
val rule'' =  | 
| 
 
6799b38ed872
added classical_rule, which replaces Data.make_elim;
 
wenzelm 
parents: 
18374 
diff
changeset
 | 
162  | 
rule' |> ALLGOALS (SUBGOAL (fn (goal, i) =>  | 
| 
 
6799b38ed872
added classical_rule, which replaces Data.make_elim;
 
wenzelm 
parents: 
18374 
diff
changeset
 | 
163  | 
if i = 1 orelse redundant_hyp goal  | 
| 60757 | 164  | 
then eresolve_tac ctxt [thin_rl] i  | 
| 
18534
 
6799b38ed872
added classical_rule, which replaces Data.make_elim;
 
wenzelm 
parents: 
18374 
diff
changeset
 | 
165  | 
else all_tac))  | 
| 
 
6799b38ed872
added classical_rule, which replaces Data.make_elim;
 
wenzelm 
parents: 
18374 
diff
changeset
 | 
166  | 
|> Seq.hd  | 
| 21963 | 167  | 
|> Drule.zero_var_indexes;  | 
| 60817 | 168  | 
in if Thm.equiv_thm thy (rule, rule'') then rule else rule'' end  | 
| 
18534
 
6799b38ed872
added classical_rule, which replaces Data.make_elim;
 
wenzelm 
parents: 
18374 
diff
changeset
 | 
169  | 
else rule;  | 
| 
 
6799b38ed872
added classical_rule, which replaces Data.make_elim;
 
wenzelm 
parents: 
18374 
diff
changeset
 | 
170  | 
|
| 82852 | 171  | 
|
172  | 
(* flatten nested meta connectives in prems *)  | 
|
173  | 
||
| 59970 | 174  | 
fun flat_rule ctxt =  | 
175  | 
Conv.fconv_rule (Conv.prems_conv ~1 (Object_Logic.atomize_prems ctxt));  | 
|
| 
18534
 
6799b38ed872
added classical_rule, which replaces Data.make_elim;
 
wenzelm 
parents: 
18374 
diff
changeset
 | 
176  | 
|
| 
 
6799b38ed872
added classical_rule, which replaces Data.make_elim;
 
wenzelm 
parents: 
18374 
diff
changeset
 | 
177  | 
|
| 82852 | 178  | 
(* Useful tactics for classical reasoning *)  | 
| 0 | 179  | 
|
| 82852 | 180  | 
(*Prove goal that assumes both P and \<not> P.  | 
| 4392 | 181  | 
No backtracking if it finds an equal assumption. Perhaps should call  | 
182  | 
ematch_tac instead of eresolve_tac, but then cannot prove ZF/cantor.*)  | 
|
| 
58963
 
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
 
wenzelm 
parents: 
58958 
diff
changeset
 | 
183  | 
fun contr_tac ctxt =  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59164 
diff
changeset
 | 
184  | 
eresolve_tac ctxt [Data.not_elim] THEN' (eq_assume_tac ORELSE' assume_tac ctxt);  | 
| 0 | 185  | 
|
| 82852 | 186  | 
(*Finds P \<longrightarrow> Q and P in the assumptions, replaces implication by Q.  | 
187  | 
Could do the same thing for P \<longleftrightarrow> Q and P.*)  | 
|
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59164 
diff
changeset
 | 
188  | 
fun mp_tac ctxt i =  | 
| 
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59164 
diff
changeset
 | 
189  | 
eresolve_tac ctxt [Data.not_elim, Data.imp_elim] i THEN assume_tac ctxt i;  | 
| 0 | 190  | 
|
191  | 
(*Like mp_tac but instantiates no variables*)  | 
|
| 58957 | 192  | 
fun eq_mp_tac ctxt i = ematch_tac ctxt [Data.not_elim, Data.imp_elim] i THEN eq_assume_tac i;  | 
| 0 | 193  | 
|
| 82852 | 194  | 
(*Creates rules to eliminate \<not> A, from rules to introduce A*)  | 
| 
82832
 
bf1bc2932343
more robust: unique result expected, otherwise index calculations will go wrong;
 
wenzelm 
parents: 
82831 
diff
changeset
 | 
195  | 
fun maybe_swap_rule th =  | 
| 
82833
 
13a8c49a48a0
clarified signature: more explicit types, notably (thm option) instead of (thm list);
 
wenzelm 
parents: 
82832 
diff
changeset
 | 
196  | 
(case [th] RLN (2, [Data.swap]) of  | 
| 
 
13a8c49a48a0
clarified signature: more explicit types, notably (thm option) instead of (thm list);
 
wenzelm 
parents: 
82832 
diff
changeset
 | 
197  | 
[] => NONE  | 
| 
 
13a8c49a48a0
clarified signature: more explicit types, notably (thm option) instead of (thm list);
 
wenzelm 
parents: 
82832 
diff
changeset
 | 
198  | 
| [res] => SOME res  | 
| 
 
13a8c49a48a0
clarified signature: more explicit types, notably (thm option) instead of (thm list);
 
wenzelm 
parents: 
82832 
diff
changeset
 | 
199  | 
  | _ => raise THM ("RSN: multiple unifiers", 2, [th, Data.swap]));
 | 
| 
82832
 
bf1bc2932343
more robust: unique result expected, otherwise index calculations will go wrong;
 
wenzelm 
parents: 
82831 
diff
changeset
 | 
200  | 
|
| 82831 | 201  | 
fun swap_rule intr = intr RSN (2, Data.swap);  | 
202  | 
val swapped = Thm.rule_attribute [] (K swap_rule);  | 
|
| 0 | 203  | 
|
204  | 
(*Uses introduction rules in the normal way, or on negated assumptions,  | 
|
205  | 
trying rules in order. *)  | 
|
| 
58963
 
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
 
wenzelm 
parents: 
58958 
diff
changeset
 | 
206  | 
fun swap_res_tac ctxt rls =  | 
| 61056 | 207  | 
let  | 
| 82831 | 208  | 
fun addrl rl brls = (false, rl) :: (true, swap_rule rl) :: brls;  | 
| 61056 | 209  | 
in  | 
| 
58963
 
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
 
wenzelm 
parents: 
58958 
diff
changeset
 | 
210  | 
assume_tac ctxt ORELSE'  | 
| 
 
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
 
wenzelm 
parents: 
58958 
diff
changeset
 | 
211  | 
contr_tac ctxt ORELSE'  | 
| 82830 | 212  | 
biresolve_tac ctxt (fold_rev (addrl o Thm.transfer' ctxt) rls [])  | 
| 42792 | 213  | 
end;  | 
| 0 | 214  | 
|
| 60943 | 215  | 
(*Duplication of unsafe rules, for complete provers*)  | 
| 42792 | 216  | 
fun dup_intr th = zero_var_indexes (th RS Data.classical);  | 
| 
681
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
217  | 
|
| 59970 | 218  | 
fun dup_elim ctxt th =  | 
219  | 
let val rl = (th RSN (2, revcut_rl)) |> Thm.assumption (SOME ctxt) 2 |> Seq.hd;  | 
|
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59164 
diff
changeset
 | 
220  | 
in rule_by_tactic ctxt (TRYALL (eresolve_tac ctxt [revcut_rl])) rl end;  | 
| 36546 | 221  | 
|
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
222  | 
|
| 82852 | 223  | 
|
224  | 
(** Classical rule sets **)  | 
|
| 0 | 225  | 
|
| 
82833
 
13a8c49a48a0
clarified signature: more explicit types, notably (thm option) instead of (thm list);
 
wenzelm 
parents: 
82832 
diff
changeset
 | 
226  | 
type rl = thm * thm option; (*internal form, with possibly swapped version*)  | 
| 
82851
 
7f9c4466c6a5
proper "plain" rule for extra_netpair (amending 8aa1c98b948b): need to avoid flat_rule / Object_Logic.atomize_prems for the sake of "standard" Isar proof, e.g. (line 34 of "$AFP/AWN/Qmsg_Lifting.thy);
 
wenzelm 
parents: 
82849 
diff
changeset
 | 
227  | 
type info = {rl: rl, dup_rl: rl, plain: thm};
 | 
| 
82833
 
13a8c49a48a0
clarified signature: more explicit types, notably (thm option) instead of (thm list);
 
wenzelm 
parents: 
82832 
diff
changeset
 | 
228  | 
type rule = thm * info; (*external form, internal forms*)  | 
| 
 
13a8c49a48a0
clarified signature: more explicit types, notably (thm option) instead of (thm list);
 
wenzelm 
parents: 
82832 
diff
changeset
 | 
229  | 
|
| 
82842
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
230  | 
type decl = info Bires.decl;  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
231  | 
type decls = info Bires.decls;  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
232  | 
|
| 
82833
 
13a8c49a48a0
clarified signature: more explicit types, notably (thm option) instead of (thm list);
 
wenzelm 
parents: 
82832 
diff
changeset
 | 
233  | 
fun maybe_swapped_rl th : rl = (th, maybe_swap_rule th);  | 
| 
 
13a8c49a48a0
clarified signature: more explicit types, notably (thm option) instead of (thm list);
 
wenzelm 
parents: 
82832 
diff
changeset
 | 
234  | 
fun no_swapped_rl th : rl = (th, NONE);  | 
| 
 
13a8c49a48a0
clarified signature: more explicit types, notably (thm option) instead of (thm list);
 
wenzelm 
parents: 
82832 
diff
changeset
 | 
235  | 
|
| 
82851
 
7f9c4466c6a5
proper "plain" rule for extra_netpair (amending 8aa1c98b948b): need to avoid flat_rule / Object_Logic.atomize_prems for the sake of "standard" Isar proof, e.g. (line 34 of "$AFP/AWN/Qmsg_Lifting.thy);
 
wenzelm 
parents: 
82849 
diff
changeset
 | 
236  | 
fun make_info rl dup_rl plain : info = {rl = rl, dup_rl = dup_rl, plain = plain};
 | 
| 
 
7f9c4466c6a5
proper "plain" rule for extra_netpair (amending 8aa1c98b948b): need to avoid flat_rule / Object_Logic.atomize_prems for the sake of "standard" Isar proof, e.g. (line 34 of "$AFP/AWN/Qmsg_Lifting.thy);
 
wenzelm 
parents: 
82849 
diff
changeset
 | 
237  | 
fun make_info1 rl plain : info = make_info rl rl plain;  | 
| 
61055
 
6fc78876f9be
store result of swapify, to avoid later access to implicit context;
 
wenzelm 
parents: 
61049 
diff
changeset
 | 
238  | 
|
| 42812 | 239  | 
type wrapper = (int -> tactic) -> int -> tactic;  | 
240  | 
||
| 0 | 241  | 
datatype claset =  | 
| 42793 | 242  | 
CS of  | 
| 
82842
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
243  | 
   {decls: decls,  (*rule declarations in canonical order*)
 | 
| 60945 | 244  | 
swrappers: (string * (Proof.context -> wrapper)) list, (*for transforming safe_step_tac*)  | 
245  | 
uwrappers: (string * (Proof.context -> wrapper)) list, (*for transforming step_tac*)  | 
|
| 
82812
 
ea8d633fd4a8
just one type Bires.netpair, based on Bires.tag with explicit weight;
 
wenzelm 
parents: 
82809 
diff
changeset
 | 
246  | 
safe0_netpair: Bires.netpair, (*nets for trivial cases*)  | 
| 
 
ea8d633fd4a8
just one type Bires.netpair, based on Bires.tag with explicit weight;
 
wenzelm 
parents: 
82809 
diff
changeset
 | 
247  | 
safep_netpair: Bires.netpair, (*nets for >0 subgoals*)  | 
| 
 
ea8d633fd4a8
just one type Bires.netpair, based on Bires.tag with explicit weight;
 
wenzelm 
parents: 
82809 
diff
changeset
 | 
248  | 
unsafe_netpair: Bires.netpair, (*nets for unsafe rules*)  | 
| 
 
ea8d633fd4a8
just one type Bires.netpair, based on Bires.tag with explicit weight;
 
wenzelm 
parents: 
82809 
diff
changeset
 | 
249  | 
dup_netpair: Bires.netpair, (*nets for duplication*)  | 
| 
 
ea8d633fd4a8
just one type Bires.netpair, based on Bires.tag with explicit weight;
 
wenzelm 
parents: 
82809 
diff
changeset
 | 
250  | 
extra_netpair: Bires.netpair}; (*nets for extra rules*)  | 
| 60945 | 251  | 
|
| 10736 | 252  | 
val empty_cs =  | 
| 42793 | 253  | 
CS  | 
| 
82842
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
254  | 
   {decls = Bires.empty_decls,
 | 
| 42793 | 255  | 
swrappers = [],  | 
256  | 
uwrappers = [],  | 
|
| 
82812
 
ea8d633fd4a8
just one type Bires.netpair, based on Bires.tag with explicit weight;
 
wenzelm 
parents: 
82809 
diff
changeset
 | 
257  | 
safe0_netpair = Bires.empty_netpair,  | 
| 
 
ea8d633fd4a8
just one type Bires.netpair, based on Bires.tag with explicit weight;
 
wenzelm 
parents: 
82809 
diff
changeset
 | 
258  | 
safep_netpair = Bires.empty_netpair,  | 
| 
 
ea8d633fd4a8
just one type Bires.netpair, based on Bires.tag with explicit weight;
 
wenzelm 
parents: 
82809 
diff
changeset
 | 
259  | 
unsafe_netpair = Bires.empty_netpair,  | 
| 
 
ea8d633fd4a8
just one type Bires.netpair, based on Bires.tag with explicit weight;
 
wenzelm 
parents: 
82809 
diff
changeset
 | 
260  | 
dup_netpair = Bires.empty_netpair,  | 
| 
 
ea8d633fd4a8
just one type Bires.netpair, based on Bires.tag with explicit weight;
 
wenzelm 
parents: 
82809 
diff
changeset
 | 
261  | 
extra_netpair = Bires.empty_netpair};  | 
| 0 | 262  | 
|
| 4653 | 263  | 
fun rep_cs (CS args) = args;  | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
264  | 
|
| 
4079
 
9df5e4f22d96
new implicit claset mechanism based on Sign.sg anytype data;
 
wenzelm 
parents: 
4066 
diff
changeset
 | 
265  | 
|
| 82853 | 266  | 
(* netpair primitives to insert / delete rules *)  | 
| 0 | 267  | 
|
| 
82842
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
268  | 
fun insert_brl i brl =  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
269  | 
  Bires.insert_tagged_rule ({weight = Bires.subgoals_of brl, index = i}, brl);
 | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
270  | 
|
| 
82842
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
271  | 
fun insert_rl kind k ((th, swapped): rl) =  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
272  | 
insert_brl (2 * k + 1) (Bires.kind_rule kind th) #>  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
273  | 
(case swapped of NONE => I | SOME th' => insert_brl (2 * k) (true, th'));  | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
274  | 
|
| 
82848
 
0a8977dcb21a
more accurate delete operation via authentic index --- minor performance tuning;
 
wenzelm 
parents: 
82846 
diff
changeset
 | 
275  | 
fun delete_rl k ((th, swapped): rl) =  | 
| 
 
0a8977dcb21a
more accurate delete operation via authentic index --- minor performance tuning;
 
wenzelm 
parents: 
82846 
diff
changeset
 | 
276  | 
Bires.delete_tagged_rule (2 * k + 1, th) #>  | 
| 
 
0a8977dcb21a
more accurate delete operation via authentic index --- minor performance tuning;
 
wenzelm 
parents: 
82846 
diff
changeset
 | 
277  | 
(case swapped of NONE => I | SOME th' => Bires.delete_tagged_rule (2 * k, th'));  | 
| 10736 | 278  | 
|
| 
82851
 
7f9c4466c6a5
proper "plain" rule for extra_netpair (amending 8aa1c98b948b): need to avoid flat_rule / Object_Logic.atomize_prems for the sake of "standard" Isar proof, e.g. (line 34 of "$AFP/AWN/Qmsg_Lifting.thy);
 
wenzelm 
parents: 
82849 
diff
changeset
 | 
279  | 
fun insert_plain_rule ({kind, tag, info = {plain = th, ...}}: decl) =
 | 
| 
 
7f9c4466c6a5
proper "plain" rule for extra_netpair (amending 8aa1c98b948b): need to avoid flat_rule / Object_Logic.atomize_prems for the sake of "standard" Isar proof, e.g. (line 34 of "$AFP/AWN/Qmsg_Lifting.thy);
 
wenzelm 
parents: 
82849 
diff
changeset
 | 
280  | 
Bires.insert_tagged_rule (tag, (Bires.kind_rule kind th));  | 
| 82814 | 281  | 
|
| 
82851
 
7f9c4466c6a5
proper "plain" rule for extra_netpair (amending 8aa1c98b948b): need to avoid flat_rule / Object_Logic.atomize_prems for the sake of "standard" Isar proof, e.g. (line 34 of "$AFP/AWN/Qmsg_Lifting.thy);
 
wenzelm 
parents: 
82849 
diff
changeset
 | 
282  | 
fun delete_plain_rule ({tag = {index, ...}, info = {plain = th, ...}, ...}: decl) =
 | 
| 
82848
 
0a8977dcb21a
more accurate delete operation via authentic index --- minor performance tuning;
 
wenzelm 
parents: 
82846 
diff
changeset
 | 
283  | 
Bires.delete_tagged_rule (index, th);  | 
| 1800 | 284  | 
|
| 82853 | 285  | 
|
286  | 
(* erros and warnings *)  | 
|
287  | 
||
| 82876 | 288  | 
fun err_thm ctxt msg th = error (msg () ^ "\n" ^ Thm.string_of_thm ctxt th);  | 
| 42793 | 289  | 
|
| 82863 | 290  | 
fun err_thm_illformed ctxt kind th =  | 
| 82876 | 291  | 
err_thm ctxt (fn () => "Ill-formed " ^ Bires.kind_name kind) th;  | 
| 82863 | 292  | 
|
| 
82842
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
293  | 
fun init_decl kind opt_weight info : decl =  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
294  | 
let val weight = the_default (Bires.kind_index kind) opt_weight  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
295  | 
  in {kind = kind, tag = Bires.weight_tag weight, info = info} end;
 | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
296  | 
|
| 60945 | 297  | 
fun warn_thm ctxt msg th =  | 
298  | 
if Context_Position.is_really_visible ctxt  | 
|
| 82876 | 299  | 
then warning (msg () ^ "\n" ^ Thm.string_of_thm ctxt th) else ();  | 
| 42793 | 300  | 
|
| 82878 | 301  | 
fun warn_kind ctxt prefix th kind =  | 
302  | 
if Context_Position.is_really_visible ctxt then  | 
|
303  | 
warn_thm ctxt (fn () => prefix ^ " " ^ Bires.kind_description kind) th  | 
|
304  | 
else ();  | 
|
| 
42807
 
e639d91d9073
more precise warnings: observe context visibility;
 
wenzelm 
parents: 
42799 
diff
changeset
 | 
305  | 
|
| 
82842
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
306  | 
fun warn_other_kinds ctxt decls th =  | 
| 82878 | 307  | 
if Context_Position.is_really_visible ctxt then  | 
308  | 
(case Bires.get_decls decls th of  | 
|
309  | 
[] => ()  | 
|
310  | 
| ds =>  | 
|
311  | 
Bires.kind_domain  | 
|
312  | 
|> filter (member (op =) (map #kind ds))  | 
|
313  | 
|> List.app (warn_kind ctxt "Rule already declared as" th))  | 
|
314  | 
else ();  | 
|
| 
1927
 
6f97cb16e453
New classical reasoner: warns of, and ignores, redundant rules.
 
paulson 
parents: 
1814 
diff
changeset
 | 
315  | 
|
| 
12376
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
316  | 
|
| 82852 | 317  | 
(* extend and merge rules *)  | 
| 
82842
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
318  | 
|
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
319  | 
local  | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
320  | 
|
| 
82842
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
321  | 
type netpairs = (Bires.netpair * Bires.netpair) * (Bires.netpair * Bires.netpair);  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
322  | 
|
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
323  | 
fun add_safe_rule (decl: decl) (netpairs: netpairs) =  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
324  | 
let  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
325  | 
    val {kind, tag = {index, ...}, info = {rl, ...}} = decl;
 | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
326  | 
val no_subgoals = Bires.no_subgoals (Bires.kind_rule kind (#1 rl));  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
327  | 
in (apfst o (if no_subgoals then apfst else apsnd)) (insert_rl kind index rl) netpairs end;  | 
| 0 | 328  | 
|
| 
82842
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
329  | 
fun add_unsafe_rule (decl: decl) ((safe_netpairs, (unsafe_netpair, dup_netpair)): netpairs) =  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
330  | 
let  | 
| 
82851
 
7f9c4466c6a5
proper "plain" rule for extra_netpair (amending 8aa1c98b948b): need to avoid flat_rule / Object_Logic.atomize_prems for the sake of "standard" Isar proof, e.g. (line 34 of "$AFP/AWN/Qmsg_Lifting.thy);
 
wenzelm 
parents: 
82849 
diff
changeset
 | 
331  | 
    val {kind, tag = {index, ...}, info = {rl, dup_rl, ...}} = decl;
 | 
| 
82842
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
332  | 
val unsafe_netpair' = insert_rl kind index rl unsafe_netpair;  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
333  | 
val dup_netpair' = insert_rl kind index dup_rl dup_netpair;  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
334  | 
in (safe_netpairs, (unsafe_netpair', dup_netpair')) end;  | 
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
335  | 
|
| 
82842
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
336  | 
fun add_rule (decl as {kind, ...}: decl) =
 | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
337  | 
if Bires.kind_safe kind then add_safe_rule decl  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
338  | 
else if Bires.kind_unsafe kind then add_unsafe_rule decl  | 
| 82887 | 339  | 
else raise Fail "Bad rule kind";  | 
| 
82842
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
340  | 
|
| 0 | 341  | 
|
| 
82833
 
13a8c49a48a0
clarified signature: more explicit types, notably (thm option) instead of (thm list);
 
wenzelm 
parents: 
82832 
diff
changeset
 | 
342  | 
fun trim_context_rl (th1, opt_th2) =  | 
| 
 
13a8c49a48a0
clarified signature: more explicit types, notably (thm option) instead of (thm list);
 
wenzelm 
parents: 
82832 
diff
changeset
 | 
343  | 
(Thm.trim_context th1, Option.map Thm.trim_context opt_th2);  | 
| 
 
13a8c49a48a0
clarified signature: more explicit types, notably (thm option) instead of (thm list);
 
wenzelm 
parents: 
82832 
diff
changeset
 | 
344  | 
|
| 
82851
 
7f9c4466c6a5
proper "plain" rule for extra_netpair (amending 8aa1c98b948b): need to avoid flat_rule / Object_Logic.atomize_prems for the sake of "standard" Isar proof, e.g. (line 34 of "$AFP/AWN/Qmsg_Lifting.thy);
 
wenzelm 
parents: 
82849 
diff
changeset
 | 
345  | 
fun trim_context_info {rl, dup_rl, plain} : info =
 | 
| 
 
7f9c4466c6a5
proper "plain" rule for extra_netpair (amending 8aa1c98b948b): need to avoid flat_rule / Object_Logic.atomize_prems for the sake of "standard" Isar proof, e.g. (line 34 of "$AFP/AWN/Qmsg_Lifting.thy);
 
wenzelm 
parents: 
82849 
diff
changeset
 | 
346  | 
  {rl = trim_context_rl rl, dup_rl = trim_context_rl dup_rl, plain = Thm.trim_context plain};
 | 
| 61056 | 347  | 
|
| 82865 | 348  | 
fun ext_info ctxt kind th =  | 
349  | 
if kind = Bires.intro_bang_kind then  | 
|
350  | 
make_info1 (maybe_swapped_rl (flat_rule ctxt th)) th  | 
|
| 
82868
 
c2a88a1cd07d
explicit "dest" rules: no longer declare [elim_format, elim];
 
wenzelm 
parents: 
82866 
diff
changeset
 | 
351  | 
else if kind = Bires.elim_bang_kind orelse kind = Bires.dest_bang_kind then  | 
| 
 
c2a88a1cd07d
explicit "dest" rules: no longer declare [elim_format, elim];
 
wenzelm 
parents: 
82866 
diff
changeset
 | 
352  | 
let  | 
| 
 
c2a88a1cd07d
explicit "dest" rules: no longer declare [elim_format, elim];
 
wenzelm 
parents: 
82866 
diff
changeset
 | 
353  | 
val _ = Thm.no_prems th andalso err_thm_illformed ctxt kind th;  | 
| 82873 | 354  | 
val elim = Bires.kind_make_elim kind th;  | 
| 
82868
 
c2a88a1cd07d
explicit "dest" rules: no longer declare [elim_format, elim];
 
wenzelm 
parents: 
82866 
diff
changeset
 | 
355  | 
in make_info1 (no_swapped_rl (classical_rule ctxt (flat_rule ctxt elim))) elim end  | 
| 82865 | 356  | 
else if kind = Bires.intro_kind then  | 
357  | 
let  | 
|
358  | 
val th' = flat_rule ctxt th;  | 
|
359  | 
val dup_th' = dup_intr th' handle THM _ => err_thm_illformed ctxt kind th;  | 
|
360  | 
in make_info (maybe_swapped_rl th') (maybe_swapped_rl dup_th') th end  | 
|
| 
82868
 
c2a88a1cd07d
explicit "dest" rules: no longer declare [elim_format, elim];
 
wenzelm 
parents: 
82866 
diff
changeset
 | 
361  | 
else if kind = Bires.elim_kind orelse kind = Bires.dest_kind then  | 
| 82865 | 362  | 
let  | 
363  | 
val _ = Thm.no_prems th andalso err_thm_illformed ctxt kind th;  | 
|
| 82873 | 364  | 
val elim = Bires.kind_make_elim kind th;  | 
| 
82868
 
c2a88a1cd07d
explicit "dest" rules: no longer declare [elim_format, elim];
 
wenzelm 
parents: 
82866 
diff
changeset
 | 
365  | 
val th' = classical_rule ctxt (flat_rule ctxt elim);  | 
| 
82866
 
bfd8258133a1
avoid redundant argument (amending af6f55b15749);
 
wenzelm 
parents: 
82865 
diff
changeset
 | 
366  | 
val dup_th' = dup_elim ctxt th' handle THM _ => err_thm_illformed ctxt kind th;  | 
| 
82868
 
c2a88a1cd07d
explicit "dest" rules: no longer declare [elim_format, elim];
 
wenzelm 
parents: 
82866 
diff
changeset
 | 
367  | 
in make_info (no_swapped_rl th') (no_swapped_rl dup_th') elim end  | 
| 82865 | 368  | 
else raise Fail "Bad rule kind";  | 
369  | 
||
370  | 
in  | 
|
371  | 
||
372  | 
fun extend_rules ctxt kind opt_weight th cs =  | 
|
| 60946 | 373  | 
let  | 
| 
82842
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
374  | 
val th' = Thm.trim_context th;  | 
| 82865 | 375  | 
val decl' = init_decl kind opt_weight (trim_context_info (ext_info ctxt kind th));  | 
| 
82842
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
376  | 
    val CS {decls, swrappers, uwrappers, safe0_netpair, safep_netpair,
 | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
377  | 
unsafe_netpair, dup_netpair, extra_netpair} = cs;  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
378  | 
in  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
379  | 
(case Bires.extend_decls (th', decl') decls of  | 
| 82878 | 380  | 
(NONE, _) => (warn_kind ctxt "Ignoring duplicate" th kind; cs)  | 
| 
82842
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
381  | 
| (SOME new_decl, decls') =>  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
382  | 
let  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
383  | 
val _ = warn_other_kinds ctxt decls th;  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
384  | 
val ((safe0_netpair', safep_netpair'), (unsafe_netpair', dup_netpair')) =  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
385  | 
add_rule new_decl ((safe0_netpair, safep_netpair), (unsafe_netpair, dup_netpair));  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
386  | 
val extra_netpair' = insert_plain_rule new_decl extra_netpair;  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
387  | 
in  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
388  | 
          CS {
 | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
389  | 
decls = decls',  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
390  | 
swrappers = swrappers,  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
391  | 
uwrappers = uwrappers,  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
392  | 
safe0_netpair = safe0_netpair',  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
393  | 
safep_netpair = safep_netpair',  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
394  | 
unsafe_netpair = unsafe_netpair',  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
395  | 
dup_netpair = dup_netpair',  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
396  | 
extra_netpair = extra_netpair'}  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
397  | 
end)  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
398  | 
end;  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
399  | 
|
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
400  | 
fun merge_cs (cs, cs2) =  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
401  | 
if pointer_eq (cs, cs2) then cs  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
402  | 
else  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
403  | 
let  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
404  | 
      val CS {decls, swrappers, uwrappers, safe0_netpair, safep_netpair,
 | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
405  | 
unsafe_netpair, dup_netpair, extra_netpair} = cs;  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
406  | 
      val CS {decls = decls2, swrappers = swrappers2, uwrappers = uwrappers2, ...} = cs2;
 | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
407  | 
|
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
408  | 
val (new_decls, decls') = Bires.merge_decls (decls, decls2);  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
409  | 
val ((safe0_netpair', safep_netpair'), (unsafe_netpair', dup_netpair')) =  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
410  | 
fold add_rule new_decls ((safe0_netpair, safep_netpair), (unsafe_netpair, dup_netpair));  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
411  | 
val extra_netpair' = fold insert_plain_rule new_decls extra_netpair;  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
412  | 
in  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
413  | 
      CS {
 | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
414  | 
decls = decls',  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
415  | 
swrappers = AList.merge (op =) (K true) (swrappers, swrappers2),  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
416  | 
uwrappers = AList.merge (op =) (K true) (uwrappers, uwrappers2),  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
417  | 
safe0_netpair = safe0_netpair',  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
418  | 
safep_netpair = safep_netpair',  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
419  | 
unsafe_netpair = unsafe_netpair',  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
420  | 
dup_netpair = dup_netpair',  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
421  | 
extra_netpair = extra_netpair'}  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
422  | 
end;  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
423  | 
|
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
424  | 
end;  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
425  | 
|
| 
1073
 
b3f190995bc9
Recoded addSIs, etc., so that nets are built incrementally
 
lcp 
parents: 
1010 
diff
changeset
 | 
426  | 
|
| 82852 | 427  | 
(* delete rules *)  | 
| 60946 | 428  | 
|
| 82874 | 429  | 
fun delrule ctxt warn th cs =  | 
| 60946 | 430  | 
let  | 
| 
82842
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
431  | 
    val CS {decls, swrappers, uwrappers,
 | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
432  | 
safe0_netpair, safep_netpair, unsafe_netpair, dup_netpair, extra_netpair} = cs;  | 
| 
82868
 
c2a88a1cd07d
explicit "dest" rules: no longer declare [elim_format, elim];
 
wenzelm 
parents: 
82866 
diff
changeset
 | 
433  | 
val (old_decls, decls') = Bires.remove_decls th decls;  | 
| 82875 | 434  | 
val _ =  | 
435  | 
if Bires.has_decls decls (Tactic.make_elim th) then  | 
|
436  | 
warn_thm ctxt  | 
|
| 82876 | 437  | 
(fn () => "Not deleting elim format --- need to to declare proper dest rule") th  | 
| 82875 | 438  | 
else ();  | 
| 60946 | 439  | 
in  | 
| 82874 | 440  | 
if null old_decls then  | 
| 82876 | 441  | 
(if warn then warn_thm ctxt (fn () => "Undeclared classical rule") th else (); cs)  | 
| 
82842
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
442  | 
else  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
443  | 
let  | 
| 
82848
 
0a8977dcb21a
more accurate delete operation via authentic index --- minor performance tuning;
 
wenzelm 
parents: 
82846 
diff
changeset
 | 
444  | 
        fun del which ({tag = {index, ...}, info, ...}: decl) = delete_rl index (which info);
 | 
| 
82842
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
445  | 
val safe0_netpair' = fold (del #rl) old_decls safe0_netpair;  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
446  | 
val safep_netpair' = fold (del #rl) old_decls safep_netpair;  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
447  | 
val unsafe_netpair' = fold (del #rl) old_decls unsafe_netpair;  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
448  | 
val dup_netpair' = fold (del #dup_rl) old_decls dup_netpair;  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
449  | 
val extra_netpair' = fold delete_plain_rule old_decls extra_netpair;  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
450  | 
in  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
451  | 
        CS {
 | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
452  | 
decls = decls',  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
453  | 
swrappers = swrappers,  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
454  | 
uwrappers = uwrappers,  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
455  | 
safe0_netpair = safe0_netpair',  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
456  | 
safep_netpair = safep_netpair',  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
457  | 
unsafe_netpair = unsafe_netpair',  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
458  | 
dup_netpair = dup_netpair',  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
459  | 
extra_netpair = extra_netpair'}  | 
| 
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
460  | 
end  | 
| 60946 | 461  | 
end;  | 
| 1800 | 462  | 
|
463  | 
||
| 82852 | 464  | 
(* Claset context data *)  | 
| 42793 | 465  | 
|
466  | 
structure Claset = Generic_Data  | 
|
467  | 
(  | 
|
468  | 
type T = claset;  | 
|
469  | 
val empty = empty_cs;  | 
|
470  | 
val merge = merge_cs;  | 
|
471  | 
);  | 
|
472  | 
||
473  | 
val claset_of = Claset.get o Context.Proof;  | 
|
474  | 
val rep_claset_of = rep_cs o claset_of;  | 
|
475  | 
||
| 82845 | 476  | 
val dest_decls = Bires.dest_decls o #decls o rep_claset_of;  | 
| 
82842
 
8aa1c98b948b
maintain collective rule declarations via type Bires.decls, with netpair operations derived from it;
 
wenzelm 
parents: 
82841 
diff
changeset
 | 
477  | 
|
| 
82846
 
c906b8df7bc3
clarified signature: do not expose internal data structures;
 
wenzelm 
parents: 
82845 
diff
changeset
 | 
478  | 
val safe0_netpair_of = #safe0_netpair o rep_claset_of;  | 
| 
 
c906b8df7bc3
clarified signature: do not expose internal data structures;
 
wenzelm 
parents: 
82845 
diff
changeset
 | 
479  | 
val safep_netpair_of = #safep_netpair o rep_claset_of;  | 
| 
 
c906b8df7bc3
clarified signature: do not expose internal data structures;
 
wenzelm 
parents: 
82845 
diff
changeset
 | 
480  | 
val unsafe_netpair_of = #unsafe_netpair o rep_claset_of;  | 
| 
 
c906b8df7bc3
clarified signature: do not expose internal data structures;
 
wenzelm 
parents: 
82845 
diff
changeset
 | 
481  | 
val dup_netpair_of = #dup_netpair o rep_claset_of;  | 
| 
 
c906b8df7bc3
clarified signature: do not expose internal data structures;
 
wenzelm 
parents: 
82845 
diff
changeset
 | 
482  | 
val extra_netpair_of = #extra_netpair o rep_claset_of;  | 
| 
 
c906b8df7bc3
clarified signature: do not expose internal data structures;
 
wenzelm 
parents: 
82845 
diff
changeset
 | 
483  | 
|
| 42793 | 484  | 
val get_cs = Claset.get;  | 
485  | 
val map_cs = Claset.map;  | 
|
486  | 
||
| 
51703
 
f2e92fc0c8aa
modifiers for classical wrappers operate on Proof.context instead of claset;
 
wenzelm 
parents: 
51688 
diff
changeset
 | 
487  | 
fun map_theory_claset f thy =  | 
| 
 
f2e92fc0c8aa
modifiers for classical wrappers operate on Proof.context instead of claset;
 
wenzelm 
parents: 
51688 
diff
changeset
 | 
488  | 
let  | 
| 
 
f2e92fc0c8aa
modifiers for classical wrappers operate on Proof.context instead of claset;
 
wenzelm 
parents: 
51688 
diff
changeset
 | 
489  | 
val ctxt' = f (Proof_Context.init_global thy);  | 
| 
 
f2e92fc0c8aa
modifiers for classical wrappers operate on Proof.context instead of claset;
 
wenzelm 
parents: 
51688 
diff
changeset
 | 
490  | 
val thy' = Proof_Context.theory_of ctxt';  | 
| 
 
f2e92fc0c8aa
modifiers for classical wrappers operate on Proof.context instead of claset;
 
wenzelm 
parents: 
51688 
diff
changeset
 | 
491  | 
in Context.theory_map (Claset.map (K (claset_of ctxt'))) thy' end;  | 
| 
 
f2e92fc0c8aa
modifiers for classical wrappers operate on Proof.context instead of claset;
 
wenzelm 
parents: 
51688 
diff
changeset
 | 
492  | 
|
| 42793 | 493  | 
fun map_claset f = Context.proof_map (map_cs f);  | 
494  | 
fun put_claset cs = map_claset (K cs);  | 
|
495  | 
||
496  | 
||
| 82865 | 497  | 
(* old-style ML declarations *)  | 
| 42793 | 498  | 
|
| 
82868
 
c2a88a1cd07d
explicit "dest" rules: no longer declare [elim_format, elim];
 
wenzelm 
parents: 
82866 
diff
changeset
 | 
499  | 
fun ml_decl kind (ctxt, ths) =  | 
| 
 
c2a88a1cd07d
explicit "dest" rules: no longer declare [elim_format, elim];
 
wenzelm 
parents: 
82866 
diff
changeset
 | 
500  | 
map_claset (fold_rev (extend_rules ctxt kind NONE) ths) ctxt;  | 
| 42793 | 501  | 
|
| 
82868
 
c2a88a1cd07d
explicit "dest" rules: no longer declare [elim_format, elim];
 
wenzelm 
parents: 
82866 
diff
changeset
 | 
502  | 
val op addSIs = ml_decl Bires.intro_bang_kind;  | 
| 
 
c2a88a1cd07d
explicit "dest" rules: no longer declare [elim_format, elim];
 
wenzelm 
parents: 
82866 
diff
changeset
 | 
503  | 
val op addSEs = ml_decl Bires.elim_bang_kind;  | 
| 
 
c2a88a1cd07d
explicit "dest" rules: no longer declare [elim_format, elim];
 
wenzelm 
parents: 
82866 
diff
changeset
 | 
504  | 
val op addSDs = ml_decl Bires.dest_bang_kind;  | 
| 
 
c2a88a1cd07d
explicit "dest" rules: no longer declare [elim_format, elim];
 
wenzelm 
parents: 
82866 
diff
changeset
 | 
505  | 
val op addIs = ml_decl Bires.intro_kind;  | 
| 
 
c2a88a1cd07d
explicit "dest" rules: no longer declare [elim_format, elim];
 
wenzelm 
parents: 
82866 
diff
changeset
 | 
506  | 
val op addEs = ml_decl Bires.elim_kind;  | 
| 
 
c2a88a1cd07d
explicit "dest" rules: no longer declare [elim_format, elim];
 
wenzelm 
parents: 
82866 
diff
changeset
 | 
507  | 
val op addDs = ml_decl Bires.dest_kind;  | 
| 82865 | 508  | 
|
| 82874 | 509  | 
fun ctxt delrules ths = map_claset (fold_rev (delrule ctxt true) ths) ctxt;  | 
| 42793 | 510  | 
|
511  | 
||
512  | 
||
| 82852 | 513  | 
(** Modifying the wrapper tacticals **)  | 
| 42793 | 514  | 
|
| 82853 | 515  | 
fun map_swrappers f  | 
516  | 
  (CS {decls, swrappers, uwrappers,
 | 
|
517  | 
safe0_netpair, safep_netpair, unsafe_netpair, dup_netpair, extra_netpair}) =  | 
|
518  | 
  CS {decls = decls, swrappers = f swrappers, uwrappers = uwrappers,
 | 
|
519  | 
safe0_netpair = safe0_netpair, safep_netpair = safep_netpair,  | 
|
520  | 
unsafe_netpair = unsafe_netpair, dup_netpair = dup_netpair, extra_netpair = extra_netpair};  | 
|
521  | 
||
522  | 
fun map_uwrappers f  | 
|
523  | 
  (CS {decls, swrappers, uwrappers,
 | 
|
524  | 
safe0_netpair, safep_netpair, unsafe_netpair, dup_netpair, extra_netpair}) =  | 
|
525  | 
  CS {decls = decls, swrappers = swrappers, uwrappers = f uwrappers,
 | 
|
526  | 
safe0_netpair = safe0_netpair, safep_netpair = safep_netpair,  | 
|
527  | 
unsafe_netpair = unsafe_netpair, dup_netpair = dup_netpair, extra_netpair = extra_netpair};  | 
|
528  | 
||
| 42793 | 529  | 
fun appSWrappers ctxt = fold (fn (_, w) => w ctxt) (#swrappers (rep_claset_of ctxt));  | 
530  | 
fun appWrappers ctxt = fold (fn (_, w) => w ctxt) (#uwrappers (rep_claset_of ctxt));  | 
|
531  | 
||
532  | 
fun update_warn msg (p as (key : string, _)) xs =  | 
|
533  | 
(if AList.defined (op =) xs key then warning msg else (); AList.update (op =) p xs);  | 
|
534  | 
||
535  | 
fun delete_warn msg (key : string) xs =  | 
|
536  | 
if AList.defined (op =) xs key then AList.delete (op =) key xs  | 
|
537  | 
else (warning msg; xs);  | 
|
538  | 
||
539  | 
(*Add/replace a safe wrapper*)  | 
|
| 
51703
 
f2e92fc0c8aa
modifiers for classical wrappers operate on Proof.context instead of claset;
 
wenzelm 
parents: 
51688 
diff
changeset
 | 
540  | 
fun ctxt addSWrapper new_swrapper = ctxt |> map_claset  | 
| 
 
f2e92fc0c8aa
modifiers for classical wrappers operate on Proof.context instead of claset;
 
wenzelm 
parents: 
51688 
diff
changeset
 | 
541  | 
  (map_swrappers (update_warn ("Overwriting safe wrapper " ^ fst new_swrapper) new_swrapper));
 | 
| 42793 | 542  | 
|
543  | 
(*Add/replace an unsafe wrapper*)  | 
|
| 
51703
 
f2e92fc0c8aa
modifiers for classical wrappers operate on Proof.context instead of claset;
 
wenzelm 
parents: 
51688 
diff
changeset
 | 
544  | 
fun ctxt addWrapper new_uwrapper = ctxt |> map_claset  | 
| 
 
f2e92fc0c8aa
modifiers for classical wrappers operate on Proof.context instead of claset;
 
wenzelm 
parents: 
51688 
diff
changeset
 | 
545  | 
  (map_uwrappers (update_warn ("Overwriting unsafe wrapper " ^ fst new_uwrapper) new_uwrapper));
 | 
| 42793 | 546  | 
|
547  | 
(*Remove a safe wrapper*)  | 
|
| 
51703
 
f2e92fc0c8aa
modifiers for classical wrappers operate on Proof.context instead of claset;
 
wenzelm 
parents: 
51688 
diff
changeset
 | 
548  | 
fun ctxt delSWrapper name = ctxt |> map_claset  | 
| 
 
f2e92fc0c8aa
modifiers for classical wrappers operate on Proof.context instead of claset;
 
wenzelm 
parents: 
51688 
diff
changeset
 | 
549  | 
  (map_swrappers (delete_warn ("No such safe wrapper in claset: " ^ name) name));
 | 
| 42793 | 550  | 
|
551  | 
(*Remove an unsafe wrapper*)  | 
|
| 
51703
 
f2e92fc0c8aa
modifiers for classical wrappers operate on Proof.context instead of claset;
 
wenzelm 
parents: 
51688 
diff
changeset
 | 
552  | 
fun ctxt delWrapper name = ctxt |> map_claset  | 
| 
 
f2e92fc0c8aa
modifiers for classical wrappers operate on Proof.context instead of claset;
 
wenzelm 
parents: 
51688 
diff
changeset
 | 
553  | 
  (map_uwrappers (delete_warn ("No such unsafe wrapper in claset: " ^ name) name));
 | 
| 42793 | 554  | 
|
| 82852 | 555  | 
(*compose a safe tactic alternatively before/after safe_step_tac*)  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51703 
diff
changeset
 | 
556  | 
fun ctxt addSbefore (name, tac1) =  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51703 
diff
changeset
 | 
557  | 
ctxt addSWrapper (name, fn ctxt => fn tac2 => tac1 ctxt ORELSE' tac2);  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51703 
diff
changeset
 | 
558  | 
fun ctxt addSafter (name, tac2) =  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51703 
diff
changeset
 | 
559  | 
ctxt addSWrapper (name, fn ctxt => fn tac1 => tac1 ORELSE' tac2 ctxt);  | 
| 42793 | 560  | 
|
| 82852 | 561  | 
(*compose a tactic alternatively before/after the step tactic*)  | 
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51703 
diff
changeset
 | 
562  | 
fun ctxt addbefore (name, tac1) =  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51703 
diff
changeset
 | 
563  | 
ctxt addWrapper (name, fn ctxt => fn tac2 => tac1 ctxt APPEND' tac2);  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51703 
diff
changeset
 | 
564  | 
fun ctxt addafter (name, tac2) =  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
51703 
diff
changeset
 | 
565  | 
ctxt addWrapper (name, fn ctxt => fn tac1 => tac1 APPEND' tac2 ctxt);  | 
| 42793 | 566  | 
|
| 58957 | 567  | 
fun ctxt addD2 (name, thm) =  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59164 
diff
changeset
 | 
568  | 
ctxt addafter (name, fn ctxt' => dresolve_tac ctxt' [thm] THEN' assume_tac ctxt');  | 
| 58957 | 569  | 
fun ctxt addE2 (name, thm) =  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59164 
diff
changeset
 | 
570  | 
ctxt addafter (name, fn ctxt' => eresolve_tac ctxt' [thm] THEN' assume_tac ctxt');  | 
| 58957 | 571  | 
fun ctxt addSD2 (name, thm) =  | 
572  | 
ctxt addSafter (name, fn ctxt' => dmatch_tac ctxt' [thm] THEN' eq_assume_tac);  | 
|
573  | 
fun ctxt addSE2 (name, thm) =  | 
|
574  | 
ctxt addSafter (name, fn ctxt' => ematch_tac ctxt' [thm] THEN' eq_assume_tac);  | 
|
| 42793 | 575  | 
|
| 1711 | 576  | 
|
| 
982
 
4fe0b642b7d5
Addition of wrappers for integration with the simplifier.
 
lcp 
parents: 
747 
diff
changeset
 | 
577  | 
|
| 82852 | 578  | 
(** Simple tactics for theorem proving **)  | 
| 0 | 579  | 
|
580  | 
(*Attack subgoals using safe inferences -- matching, not resolution*)  | 
|
| 42793 | 581  | 
fun safe_step_tac ctxt =  | 
| 
82846
 
c906b8df7bc3
clarified signature: do not expose internal data structures;
 
wenzelm 
parents: 
82845 
diff
changeset
 | 
582  | 
appSWrappers ctxt  | 
| 
 
c906b8df7bc3
clarified signature: do not expose internal data structures;
 
wenzelm 
parents: 
82845 
diff
changeset
 | 
583  | 
(FIRST'  | 
| 
 
c906b8df7bc3
clarified signature: do not expose internal data structures;
 
wenzelm 
parents: 
82845 
diff
changeset
 | 
584  | 
[eq_assume_tac,  | 
| 
 
c906b8df7bc3
clarified signature: do not expose internal data structures;
 
wenzelm 
parents: 
82845 
diff
changeset
 | 
585  | 
eq_mp_tac ctxt,  | 
| 
 
c906b8df7bc3
clarified signature: do not expose internal data structures;
 
wenzelm 
parents: 
82845 
diff
changeset
 | 
586  | 
Bires.bimatch_from_nets_tac ctxt (safe0_netpair_of ctxt),  | 
| 
 
c906b8df7bc3
clarified signature: do not expose internal data structures;
 
wenzelm 
parents: 
82845 
diff
changeset
 | 
587  | 
FIRST' (map (fn tac => tac ctxt) Data.hyp_subst_tacs),  | 
| 
 
c906b8df7bc3
clarified signature: do not expose internal data structures;
 
wenzelm 
parents: 
82845 
diff
changeset
 | 
588  | 
Bires.bimatch_from_nets_tac ctxt (safep_netpair_of ctxt)]);  | 
| 0 | 589  | 
|
| 5757 | 590  | 
(*Repeatedly attack a subgoal using safe inferences -- it's deterministic!*)  | 
| 42793 | 591  | 
fun safe_steps_tac ctxt =  | 
592  | 
REPEAT_DETERM1 o (fn i => COND (has_fewer_prems i) no_tac (safe_step_tac ctxt i));  | 
|
| 5757 | 593  | 
|
| 0 | 594  | 
(*Repeatedly attack subgoals using safe inferences -- it's deterministic!*)  | 
| 42793 | 595  | 
fun safe_tac ctxt = REPEAT_DETERM1 (FIRSTGOAL (safe_steps_tac ctxt));  | 
| 
747
 
bdc066781063
deepen_tac: modified due to outcome of experiments.  Its
 
lcp 
parents: 
681 
diff
changeset
 | 
596  | 
|
| 
3705
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
597  | 
|
| 82852 | 598  | 
|
599  | 
(** Clarify_tac: do safe steps without causing branching **)  | 
|
| 
3705
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
600  | 
|
| 82805 | 601  | 
(*version of Bires.bimatch_from_nets_tac that only applies rules that  | 
| 
3705
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
602  | 
create precisely n subgoals.*)  | 
| 59164 | 603  | 
fun n_bimatch_from_nets_tac ctxt n =  | 
| 
82812
 
ea8d633fd4a8
just one type Bires.netpair, based on Bires.tag with explicit weight;
 
wenzelm 
parents: 
82809 
diff
changeset
 | 
604  | 
Bires.biresolution_from_nets_tac ctxt Bires.tag_ord  | 
| 
 
ea8d633fd4a8
just one type Bires.netpair, based on Bires.tag with explicit weight;
 
wenzelm 
parents: 
82809 
diff
changeset
 | 
605  | 
(SOME (fn rl => Bires.subgoals_of rl = n)) true;  | 
| 
3705
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
606  | 
|
| 58957 | 607  | 
fun eq_contr_tac ctxt i = ematch_tac ctxt [Data.not_elim] i THEN eq_assume_tac i;  | 
608  | 
fun eq_assume_contr_tac ctxt = eq_assume_tac ORELSE' eq_contr_tac ctxt;  | 
|
| 
3705
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
609  | 
|
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
610  | 
(*Two-way branching is allowed only if one of the branches immediately closes*)  | 
| 58957 | 611  | 
fun bimatch2_tac ctxt netpair i =  | 
| 59164 | 612  | 
n_bimatch_from_nets_tac ctxt 2 netpair i THEN  | 
| 58957 | 613  | 
(eq_assume_contr_tac ctxt i ORELSE eq_assume_contr_tac ctxt (i + 1));  | 
| 
3705
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
614  | 
|
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
615  | 
(*Attack subgoals using safe inferences -- matching, not resolution*)  | 
| 42793 | 616  | 
fun clarify_step_tac ctxt =  | 
| 
82846
 
c906b8df7bc3
clarified signature: do not expose internal data structures;
 
wenzelm 
parents: 
82845 
diff
changeset
 | 
617  | 
appSWrappers ctxt  | 
| 
 
c906b8df7bc3
clarified signature: do not expose internal data structures;
 
wenzelm 
parents: 
82845 
diff
changeset
 | 
618  | 
(FIRST'  | 
| 
 
c906b8df7bc3
clarified signature: do not expose internal data structures;
 
wenzelm 
parents: 
82845 
diff
changeset
 | 
619  | 
[eq_assume_contr_tac ctxt,  | 
| 
 
c906b8df7bc3
clarified signature: do not expose internal data structures;
 
wenzelm 
parents: 
82845 
diff
changeset
 | 
620  | 
Bires.bimatch_from_nets_tac ctxt (safe0_netpair_of ctxt),  | 
| 
 
c906b8df7bc3
clarified signature: do not expose internal data structures;
 
wenzelm 
parents: 
82845 
diff
changeset
 | 
621  | 
FIRST' (map (fn tac => tac ctxt) Data.hyp_subst_tacs),  | 
| 
 
c906b8df7bc3
clarified signature: do not expose internal data structures;
 
wenzelm 
parents: 
82845 
diff
changeset
 | 
622  | 
n_bimatch_from_nets_tac ctxt 1 (safep_netpair_of ctxt),  | 
| 
 
c906b8df7bc3
clarified signature: do not expose internal data structures;
 
wenzelm 
parents: 
82845 
diff
changeset
 | 
623  | 
bimatch2_tac ctxt (safep_netpair_of ctxt)]);  | 
| 
3705
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
624  | 
|
| 42793 | 625  | 
fun clarify_tac ctxt = SELECT_GOAL (REPEAT_DETERM (clarify_step_tac ctxt 1));  | 
| 
3705
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
626  | 
|
| 
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
627  | 
|
| 82852 | 628  | 
|
629  | 
(** Unsafe steps instantiate variables or lose information **)  | 
|
| 
3705
 
76f3b2803982
Addition of clarify_tac, clarify_step_tac, Clarify_tac, Clarify_step_tac
 
paulson 
parents: 
3546 
diff
changeset
 | 
630  | 
|
| 4066 | 631  | 
(*Backtracking is allowed among the various these unsafe ways of  | 
632  | 
proving a subgoal. *)  | 
|
| 42793 | 633  | 
fun inst0_step_tac ctxt =  | 
| 
58963
 
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
 
wenzelm 
parents: 
58958 
diff
changeset
 | 
634  | 
assume_tac ctxt APPEND'  | 
| 
 
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
 
wenzelm 
parents: 
58958 
diff
changeset
 | 
635  | 
contr_tac ctxt APPEND'  | 
| 
82846
 
c906b8df7bc3
clarified signature: do not expose internal data structures;
 
wenzelm 
parents: 
82845 
diff
changeset
 | 
636  | 
Bires.biresolve_from_nets_tac ctxt (safe0_netpair_of ctxt);  | 
| 
747
 
bdc066781063
deepen_tac: modified due to outcome of experiments.  Its
 
lcp 
parents: 
681 
diff
changeset
 | 
637  | 
|
| 4066 | 638  | 
(*These unsafe steps could generate more subgoals.*)  | 
| 42793 | 639  | 
fun instp_step_tac ctxt =  | 
| 
82846
 
c906b8df7bc3
clarified signature: do not expose internal data structures;
 
wenzelm 
parents: 
82845 
diff
changeset
 | 
640  | 
Bires.biresolve_from_nets_tac ctxt (safep_netpair_of ctxt);  | 
| 0 | 641  | 
|
642  | 
(*These steps could instantiate variables and are therefore unsafe.*)  | 
|
| 42793 | 643  | 
fun inst_step_tac ctxt = inst0_step_tac ctxt APPEND' instp_step_tac ctxt;  | 
| 0 | 644  | 
|
| 60943 | 645  | 
fun unsafe_step_tac ctxt =  | 
| 
82846
 
c906b8df7bc3
clarified signature: do not expose internal data structures;
 
wenzelm 
parents: 
82845 
diff
changeset
 | 
646  | 
Bires.biresolve_from_nets_tac ctxt (unsafe_netpair_of ctxt);  | 
| 
681
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
647  | 
|
| 0 | 648  | 
(*Single step for the prover. FAILS unless it makes progress. *)  | 
| 42793 | 649  | 
fun step_tac ctxt i =  | 
| 60943 | 650  | 
safe_tac ctxt ORELSE appWrappers ctxt (inst_step_tac ctxt ORELSE' unsafe_step_tac ctxt) i;  | 
| 0 | 651  | 
|
652  | 
(*Using a "safe" rule to instantiate variables is unsafe. This tactic  | 
|
653  | 
allows backtracking from "safe" rules to "unsafe" rules here.*)  | 
|
| 42793 | 654  | 
fun slow_step_tac ctxt i =  | 
| 60943 | 655  | 
safe_tac ctxt ORELSE appWrappers ctxt (inst_step_tac ctxt APPEND' unsafe_step_tac ctxt) i;  | 
| 0 | 656  | 
|
| 
42791
 
36f787ae5f70
eliminated weight_ASTAR: int Unsynchronized.ref (astar_tac appears to be obsolete anyway);
 
wenzelm 
parents: 
42790 
diff
changeset
 | 
657  | 
|
| 82852 | 658  | 
(** The following tactics all fail unless they solve one goal **)  | 
| 0 | 659  | 
|
660  | 
(*Dumb but fast*)  | 
|
| 42793 | 661  | 
fun fast_tac ctxt =  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
52732 
diff
changeset
 | 
662  | 
Object_Logic.atomize_prems_tac ctxt THEN' SELECT_GOAL (DEPTH_SOLVE (step_tac ctxt 1));  | 
| 0 | 663  | 
|
664  | 
(*Slower but smarter than fast_tac*)  | 
|
| 42793 | 665  | 
fun best_tac ctxt =  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
52732 
diff
changeset
 | 
666  | 
Object_Logic.atomize_prems_tac ctxt THEN'  | 
| 82804 | 667  | 
SELECT_GOAL (BEST_FIRST (Thm.no_prems, Data.sizef) (step_tac ctxt 1));  | 
| 0 | 668  | 
|
| 9402 | 669  | 
(*even a bit smarter than best_tac*)  | 
| 42793 | 670  | 
fun first_best_tac ctxt =  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
52732 
diff
changeset
 | 
671  | 
Object_Logic.atomize_prems_tac ctxt THEN'  | 
| 82804 | 672  | 
SELECT_GOAL (BEST_FIRST (Thm.no_prems, Data.sizef) (FIRSTGOAL (step_tac ctxt)));  | 
| 9402 | 673  | 
|
| 42793 | 674  | 
fun slow_tac ctxt =  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
52732 
diff
changeset
 | 
675  | 
Object_Logic.atomize_prems_tac ctxt THEN'  | 
| 42793 | 676  | 
SELECT_GOAL (DEPTH_SOLVE (slow_step_tac ctxt 1));  | 
| 0 | 677  | 
|
| 42793 | 678  | 
fun slow_best_tac ctxt =  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
52732 
diff
changeset
 | 
679  | 
Object_Logic.atomize_prems_tac ctxt THEN'  | 
| 82804 | 680  | 
SELECT_GOAL (BEST_FIRST (Thm.no_prems, Data.sizef) (slow_step_tac ctxt 1));  | 
| 0 | 681  | 
|
| 
681
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
682  | 
|
| 82852 | 683  | 
(** ASTAR with weight weight_ASTAR, by Norbert Voelker **)  | 
| 
42791
 
36f787ae5f70
eliminated weight_ASTAR: int Unsynchronized.ref (astar_tac appears to be obsolete anyway);
 
wenzelm 
parents: 
42790 
diff
changeset
 | 
684  | 
|
| 
 
36f787ae5f70
eliminated weight_ASTAR: int Unsynchronized.ref (astar_tac appears to be obsolete anyway);
 
wenzelm 
parents: 
42790 
diff
changeset
 | 
685  | 
val weight_ASTAR = 5;  | 
| 
1587
 
e7d8a4957bac
Now provides astar versions (thanks to Norbert Voelker)
 
paulson 
parents: 
1524 
diff
changeset
 | 
686  | 
|
| 42793 | 687  | 
fun astar_tac ctxt =  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
52732 
diff
changeset
 | 
688  | 
Object_Logic.atomize_prems_tac ctxt THEN'  | 
| 
10382
 
1fb807260ff1
atomize: all automated tactics that "solve" goals;
 
wenzelm 
parents: 
10309 
diff
changeset
 | 
689  | 
SELECT_GOAL  | 
| 82804 | 690  | 
(ASTAR (Thm.no_prems, fn lev => fn thm => Data.sizef thm + weight_ASTAR * lev)  | 
| 42793 | 691  | 
(step_tac ctxt 1));  | 
| 
1587
 
e7d8a4957bac
Now provides astar versions (thanks to Norbert Voelker)
 
paulson 
parents: 
1524 
diff
changeset
 | 
692  | 
|
| 42793 | 693  | 
fun slow_astar_tac ctxt =  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
52732 
diff
changeset
 | 
694  | 
Object_Logic.atomize_prems_tac ctxt THEN'  | 
| 
10382
 
1fb807260ff1
atomize: all automated tactics that "solve" goals;
 
wenzelm 
parents: 
10309 
diff
changeset
 | 
695  | 
SELECT_GOAL  | 
| 82804 | 696  | 
(ASTAR (Thm.no_prems, fn lev => fn thm => Data.sizef thm + weight_ASTAR * lev)  | 
| 42793 | 697  | 
(slow_step_tac ctxt 1));  | 
| 
1587
 
e7d8a4957bac
Now provides astar versions (thanks to Norbert Voelker)
 
paulson 
parents: 
1524 
diff
changeset
 | 
698  | 
|
| 42790 | 699  | 
|
| 82852 | 700  | 
(** Complete tactic, loosely based upon LeanTaP. This tactic is the outcome  | 
| 
747
 
bdc066781063
deepen_tac: modified due to outcome of experiments.  Its
 
lcp 
parents: 
681 
diff
changeset
 | 
701  | 
of much experimentation! Changing APPEND to ORELSE below would prove  | 
| 
 
bdc066781063
deepen_tac: modified due to outcome of experiments.  Its
 
lcp 
parents: 
681 
diff
changeset
 | 
702  | 
easy theorems faster, but loses completeness -- and many of the harder  | 
| 82852 | 703  | 
theorems such as 43. **)  | 
| 
681
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
704  | 
|
| 
747
 
bdc066781063
deepen_tac: modified due to outcome of experiments.  Its
 
lcp 
parents: 
681 
diff
changeset
 | 
705  | 
(*Non-deterministic! Could always expand the first unsafe connective.  | 
| 
 
bdc066781063
deepen_tac: modified due to outcome of experiments.  Its
 
lcp 
parents: 
681 
diff
changeset
 | 
706  | 
That's hard to implement and did not perform better in experiments, due to  | 
| 
 
bdc066781063
deepen_tac: modified due to outcome of experiments.  Its
 
lcp 
parents: 
681 
diff
changeset
 | 
707  | 
greater search depth required.*)  | 
| 42793 | 708  | 
fun dup_step_tac ctxt =  | 
| 
82846
 
c906b8df7bc3
clarified signature: do not expose internal data structures;
 
wenzelm 
parents: 
82845 
diff
changeset
 | 
709  | 
Bires.biresolve_from_nets_tac ctxt (dup_netpair_of ctxt);  | 
| 
681
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
710  | 
|
| 5523 | 711  | 
(*Searching to depth m. A variant called nodup_depth_tac appears in clasimp.ML*)  | 
| 5757 | 712  | 
local  | 
| 42793 | 713  | 
fun slow_step_tac' ctxt = appWrappers ctxt (instp_step_tac ctxt APPEND' dup_step_tac ctxt);  | 
| 42790 | 714  | 
in  | 
| 42793 | 715  | 
fun depth_tac ctxt m i state = SELECT_GOAL  | 
716  | 
(safe_steps_tac ctxt 1 THEN_ELSE  | 
|
717  | 
(DEPTH_SOLVE (depth_tac ctxt m 1),  | 
|
718  | 
inst0_step_tac ctxt 1 APPEND COND (K (m = 0)) no_tac  | 
|
719  | 
(slow_step_tac' ctxt 1 THEN DEPTH_SOLVE (depth_tac ctxt (m - 1) 1)))) i state;  | 
|
| 5757 | 720  | 
end;  | 
| 
747
 
bdc066781063
deepen_tac: modified due to outcome of experiments.  Its
 
lcp 
parents: 
681 
diff
changeset
 | 
721  | 
|
| 10736 | 722  | 
(*Search, with depth bound m.  | 
| 2173 | 723  | 
This is the "entry point", which does safe inferences first.*)  | 
| 42793 | 724  | 
fun safe_depth_tac ctxt m = SUBGOAL (fn (prem, i) =>  | 
725  | 
let  | 
|
726  | 
val deti = (*No Vars in the goal? No need to backtrack between goals.*)  | 
|
727  | 
if exists_subterm (fn Var _ => true | _ => false) prem then DETERM else I;  | 
|
| 42790 | 728  | 
in  | 
| 42793 | 729  | 
SELECT_GOAL (TRY (safe_tac ctxt) THEN DEPTH_SOLVE (deti (depth_tac ctxt m 1))) i  | 
| 42790 | 730  | 
end);  | 
| 
681
 
9b02474744ca
Provers/classical: now takes theorem "classical" as argument, proves "swap"
 
lcp 
parents: 
469 
diff
changeset
 | 
731  | 
|
| 42793 | 732  | 
fun deepen_tac ctxt = DEEPEN (2, 10) (safe_depth_tac ctxt);  | 
| 24021 | 733  | 
|
734  | 
||
| 82852 | 735  | 
|
736  | 
(** attributes **)  | 
|
737  | 
||
738  | 
(* declarations *)  | 
|
| 5885 | 739  | 
|
| 
82868
 
c2a88a1cd07d
explicit "dest" rules: no longer declare [elim_format, elim];
 
wenzelm 
parents: 
82866 
diff
changeset
 | 
740  | 
fun attrib kind w =  | 
| 60945 | 741  | 
Thm.declaration_attribute (fn th => fn context =>  | 
| 
82868
 
c2a88a1cd07d
explicit "dest" rules: no longer declare [elim_format, elim];
 
wenzelm 
parents: 
82866 
diff
changeset
 | 
742  | 
map_cs (extend_rules (Context.proof_of context) kind w th) context);  | 
| 5885 | 743  | 
|
| 
82868
 
c2a88a1cd07d
explicit "dest" rules: no longer declare [elim_format, elim];
 
wenzelm 
parents: 
82866 
diff
changeset
 | 
744  | 
val safe_intro = attrib Bires.intro_bang_kind;  | 
| 
 
c2a88a1cd07d
explicit "dest" rules: no longer declare [elim_format, elim];
 
wenzelm 
parents: 
82866 
diff
changeset
 | 
745  | 
val safe_elim = attrib Bires.elim_bang_kind;  | 
| 
 
c2a88a1cd07d
explicit "dest" rules: no longer declare [elim_format, elim];
 
wenzelm 
parents: 
82866 
diff
changeset
 | 
746  | 
val safe_dest = attrib Bires.dest_bang_kind;  | 
| 
 
c2a88a1cd07d
explicit "dest" rules: no longer declare [elim_format, elim];
 
wenzelm 
parents: 
82866 
diff
changeset
 | 
747  | 
val unsafe_intro = attrib Bires.intro_kind;  | 
| 
 
c2a88a1cd07d
explicit "dest" rules: no longer declare [elim_format, elim];
 
wenzelm 
parents: 
82866 
diff
changeset
 | 
748  | 
val unsafe_elim = attrib Bires.elim_kind;  | 
| 
 
c2a88a1cd07d
explicit "dest" rules: no longer declare [elim_format, elim];
 
wenzelm 
parents: 
82866 
diff
changeset
 | 
749  | 
val unsafe_dest = attrib Bires.dest_kind;  | 
| 
45375
 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 
wenzelm 
parents: 
42817 
diff
changeset
 | 
750  | 
|
| 
 
7fe19930dfc9
more explicit representation of rule_attribute vs. declaration_attribute vs. mixed_attribute;
 
wenzelm 
parents: 
42817 
diff
changeset
 | 
751  | 
val rule_del =  | 
| 60945 | 752  | 
Thm.declaration_attribute (fn th => fn context =>  | 
753  | 
context  | 
|
| 82874 | 754  | 
|> map_cs (delrule (Context.proof_of context) (not (Context_Rules.declared_rule context th)) th)  | 
| 60945 | 755  | 
|> Thm.attribute_declaration Context_Rules.rule_del th);  | 
| 5885 | 756  | 
|
757  | 
||
| 82852 | 758  | 
(* concrete syntax *)  | 
| 5841 | 759  | 
|
760  | 
val introN = "intro";  | 
|
761  | 
val elimN = "elim";  | 
|
762  | 
val destN = "dest";  | 
|
763  | 
||
| 58826 | 764  | 
val _ =  | 
765  | 
Theory.setup  | 
|
| 69593 | 766  | 
(Attrib.setup \<^binding>\<open>swapped\<close> (Scan.succeed swapped)  | 
| 58826 | 767  | 
"classical swap of introduction rule" #>  | 
| 69593 | 768  | 
Attrib.setup \<^binding>\<open>dest\<close> (Context_Rules.add safe_dest unsafe_dest Context_Rules.dest_query)  | 
| 58826 | 769  | 
"declaration of Classical destruction rule" #>  | 
| 69593 | 770  | 
Attrib.setup \<^binding>\<open>elim\<close> (Context_Rules.add safe_elim unsafe_elim Context_Rules.elim_query)  | 
| 58826 | 771  | 
"declaration of Classical elimination rule" #>  | 
| 69593 | 772  | 
Attrib.setup \<^binding>\<open>intro\<close> (Context_Rules.add safe_intro unsafe_intro Context_Rules.intro_query)  | 
| 58826 | 773  | 
"declaration of Classical introduction rule" #>  | 
| 69593 | 774  | 
Attrib.setup \<^binding>\<open>rule\<close> (Scan.lift Args.del >> K rule_del)  | 
| 82818 | 775  | 
"remove declaration of Classical intro/elim/dest rule");  | 
| 5841 | 776  | 
|
777  | 
||
778  | 
||
| 7230 | 779  | 
(** proof methods **)  | 
780  | 
||
781  | 
local  | 
|
782  | 
||
| 
30609
 
983e8b6e4e69
Disposed old declarations, tactics, tactic combinators that refer to the simpset or claset of an implicit theory;
 
wenzelm 
parents: 
30558 
diff
changeset
 | 
783  | 
fun some_rule_tac ctxt facts = SUBGOAL (fn (goal, i) =>  | 
| 5841 | 784  | 
let  | 
| 61049 | 785  | 
val [rules1, rules2, rules4] = Context_Rules.find_rules ctxt false facts goal;  | 
| 
82846
 
c906b8df7bc3
clarified signature: do not expose internal data structures;
 
wenzelm 
parents: 
82845 
diff
changeset
 | 
786  | 
val rules3 = Context_Rules.find_rules_netpair ctxt true facts goal (extra_netpair_of ctxt);  | 
| 
12376
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
787  | 
val rules = rules1 @ rules2 @ rules3 @ rules4;  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
58893 
diff
changeset
 | 
788  | 
val ruleq = Drule.multi_resolves (SOME ctxt) facts rules;  | 
| 52732 | 789  | 
val _ = Method.trace ctxt rules;  | 
| 
12376
 
480303e3fa08
simplified (and clarified) integration with Pure/ContextRules;
 
wenzelm 
parents: 
12362 
diff
changeset
 | 
790  | 
in  | 
| 
59498
 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 
wenzelm 
parents: 
59164 
diff
changeset
 | 
791  | 
fn st => Seq.maps (fn rule => resolve_tac ctxt [rule] i st) ruleq  | 
| 18834 | 792  | 
end)  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
52732 
diff
changeset
 | 
793  | 
THEN_ALL_NEW Goal.norm_hhf_tac ctxt;  | 
| 5841 | 794  | 
|
| 
30609
 
983e8b6e4e69
Disposed old declarations, tactics, tactic combinators that refer to the simpset or claset of an implicit theory;
 
wenzelm 
parents: 
30558 
diff
changeset
 | 
795  | 
in  | 
| 7281 | 796  | 
|
| 
30609
 
983e8b6e4e69
Disposed old declarations, tactics, tactic combinators that refer to the simpset or claset of an implicit theory;
 
wenzelm 
parents: 
30558 
diff
changeset
 | 
797  | 
fun rule_tac ctxt [] facts = some_rule_tac ctxt facts  | 
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
52732 
diff
changeset
 | 
798  | 
| rule_tac ctxt rules facts = Method.rule_tac ctxt rules facts;  | 
| 
30609
 
983e8b6e4e69
Disposed old declarations, tactics, tactic combinators that refer to the simpset or claset of an implicit theory;
 
wenzelm 
parents: 
30558 
diff
changeset
 | 
799  | 
|
| 
60619
 
7512716f03db
no arguments for "standard" (or old "default") methods;
 
wenzelm 
parents: 
60618 
diff
changeset
 | 
800  | 
fun standard_tac ctxt facts =  | 
| 
 
7512716f03db
no arguments for "standard" (or old "default") methods;
 
wenzelm 
parents: 
60618 
diff
changeset
 | 
801  | 
HEADGOAL (some_rule_tac ctxt facts) ORELSE  | 
| 
60618
 
4c79543cc376
renamed "default" to "standard", to make semantically clear what it is;
 
wenzelm 
parents: 
60097 
diff
changeset
 | 
802  | 
Class.standard_intro_classes_tac ctxt facts;  | 
| 
 
4c79543cc376
renamed "default" to "standard", to make semantically clear what it is;
 
wenzelm 
parents: 
60097 
diff
changeset
 | 
803  | 
|
| 7230 | 804  | 
end;  | 
| 5841 | 805  | 
|
806  | 
||
| 6502 | 807  | 
(* automatic methods *)  | 
| 5841 | 808  | 
|
| 5927 | 809  | 
val cla_modifiers =  | 
| 64556 | 810  | 
[Args.$$$ destN -- Args.bang_colon >> K (Method.modifier (safe_dest NONE) \<^here>),  | 
811  | 
Args.$$$ destN -- Args.colon >> K (Method.modifier (unsafe_dest NONE) \<^here>),  | 
|
812  | 
Args.$$$ elimN -- Args.bang_colon >> K (Method.modifier (safe_elim NONE) \<^here>),  | 
|
813  | 
Args.$$$ elimN -- Args.colon >> K (Method.modifier (unsafe_elim NONE) \<^here>),  | 
|
814  | 
Args.$$$ introN -- Args.bang_colon >> K (Method.modifier (safe_intro NONE) \<^here>),  | 
|
815  | 
Args.$$$ introN -- Args.colon >> K (Method.modifier (unsafe_intro NONE) \<^here>),  | 
|
816  | 
Args.del -- Args.colon >> K (Method.modifier rule_del \<^here>)];  | 
|
| 5927 | 817  | 
|
| 42793 | 818  | 
fun cla_method tac = Method.sections cla_modifiers >> K (SIMPLE_METHOD o tac);  | 
819  | 
fun cla_method' tac = Method.sections cla_modifiers >> K (SIMPLE_METHOD' o tac);  | 
|
| 5841 | 820  | 
|
821  | 
||
| 82852 | 822  | 
(* method setup *)  | 
| 5841 | 823  | 
|
| 58826 | 824  | 
val _ =  | 
825  | 
Theory.setup  | 
|
| 69593 | 826  | 
(Method.setup \<^binding>\<open>standard\<close> (Scan.succeed (METHOD o standard_tac))  | 
| 
60618
 
4c79543cc376
renamed "default" to "standard", to make semantically clear what it is;
 
wenzelm 
parents: 
60097 
diff
changeset
 | 
827  | 
"standard proof step: classical intro/elim rule or class introduction" #>  | 
| 69593 | 828  | 
Method.setup \<^binding>\<open>rule\<close>  | 
| 58826 | 829  | 
(Attrib.thms >> (fn rules => fn ctxt => METHOD (HEADGOAL o rule_tac ctxt rules)))  | 
830  | 
"apply some intro/elim rule (potentially classical)" #>  | 
|
| 69593 | 831  | 
Method.setup \<^binding>\<open>contradiction\<close>  | 
| 58826 | 832  | 
(Scan.succeed (fn ctxt => Method.rule ctxt [Data.not_elim, Drule.rotate_prems 1 Data.not_elim]))  | 
833  | 
"proof by contradiction" #>  | 
|
| 69593 | 834  | 
Method.setup \<^binding>\<open>clarify\<close> (cla_method' (CHANGED_PROP oo clarify_tac))  | 
| 58826 | 835  | 
"repeatedly apply safe steps" #>  | 
| 69593 | 836  | 
Method.setup \<^binding>\<open>fast\<close> (cla_method' fast_tac) "classical prover (depth-first)" #>  | 
837  | 
Method.setup \<^binding>\<open>slow\<close> (cla_method' slow_tac) "classical prover (slow depth-first)" #>  | 
|
838  | 
Method.setup \<^binding>\<open>best\<close> (cla_method' best_tac) "classical prover (best-first)" #>  | 
|
839  | 
Method.setup \<^binding>\<open>deepen\<close>  | 
|
| 58826 | 840  | 
(Scan.lift (Scan.optional Parse.nat 4) --| Method.sections cla_modifiers  | 
841  | 
>> (fn n => fn ctxt => SIMPLE_METHOD' (deepen_tac ctxt n)))  | 
|
842  | 
"classical prover (iterative deepening)" #>  | 
|
| 69593 | 843  | 
Method.setup \<^binding>\<open>safe\<close> (cla_method (CHANGED_PROP o safe_tac))  | 
| 58826 | 844  | 
"classical prover (apply safe rules)" #>  | 
| 69593 | 845  | 
Method.setup \<^binding>\<open>safe_step\<close> (cla_method' safe_step_tac)  | 
| 58826 | 846  | 
"single classical step (safe rules)" #>  | 
| 69593 | 847  | 
Method.setup \<^binding>\<open>inst_step\<close> (cla_method' inst_step_tac)  | 
| 58826 | 848  | 
"single classical step (safe rules, allow instantiations)" #>  | 
| 69593 | 849  | 
Method.setup \<^binding>\<open>step\<close> (cla_method' step_tac)  | 
| 58826 | 850  | 
"single classical step (safe and unsafe rules)" #>  | 
| 69593 | 851  | 
Method.setup \<^binding>\<open>slow_step\<close> (cla_method' slow_step_tac)  | 
| 58826 | 852  | 
"single classical step (safe and unsafe rules, allow backtracking)" #>  | 
| 69593 | 853  | 
Method.setup \<^binding>\<open>clarify_step\<close> (cla_method' clarify_step_tac)  | 
| 58826 | 854  | 
"single classical step (safe rules, without splitting)");  | 
| 5841 | 855  | 
|
| 8667 | 856  | 
|
| 82852 | 857  | 
|
858  | 
(** outer syntax commands **)  | 
|
| 8667 | 859  | 
|
| 82853 | 860  | 
fun print_claset ctxt =  | 
861  | 
let  | 
|
862  | 
    val {decls, swrappers, uwrappers, ...} = rep_claset_of ctxt;
 | 
|
863  | 
val prt_wrappers =  | 
|
864  | 
     [Pretty.strs ("safe wrappers:" :: map #1 swrappers),
 | 
|
865  | 
      Pretty.strs ("unsafe wrappers:" :: map #1 uwrappers)];
 | 
|
| 82864 | 866  | 
in Pretty.writeln (Pretty.chunks (Bires.pretty_decls ctxt decls @ prt_wrappers)) end;  | 
| 82853 | 867  | 
|
| 24867 | 868  | 
val _ =  | 
| 69593 | 869  | 
Outer_Syntax.command \<^command_keyword>\<open>print_claset\<close> "print context of Classical Reasoner"  | 
| 
60097
 
d20ca79d50e4
discontinued pointless warnings: commands are only defined inside a theory context;
 
wenzelm 
parents: 
59970 
diff
changeset
 | 
870  | 
(Scan.succeed (Toplevel.keep (print_claset o Toplevel.context_of)));  | 
| 8667 | 871  | 
|
| 5841 | 872  | 
end;  |