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