author | blanchet |
Wed, 04 Jan 2012 00:32:02 +0100 | |
changeset 46109 | 03e3b4b401e9 |
parent 45458 | 5b5d3ee2285b |
child 47415 | c486ac962b89 |
permissions | -rw-r--r-- |
17980 | 1 |
(* Title: Pure/goal.ML |
29345 | 2 |
Author: Makarius |
17980 | 3 |
|
18139 | 4 |
Goals in tactical theorem proving. |
17980 | 5 |
*) |
6 |
||
7 |
signature BASIC_GOAL = |
|
8 |
sig |
|
32738 | 9 |
val parallel_proofs: int Unsynchronized.ref |
41703
d27950860514
parallelization of nested Isar proofs is subject to Goal.parallel_proofs_threshold;
wenzelm
parents:
41683
diff
changeset
|
10 |
val parallel_proofs_threshold: int Unsynchronized.ref |
17980 | 11 |
val SELECT_GOAL: tactic -> int -> tactic |
23418 | 12 |
val CONJUNCTS: tactic -> int -> tactic |
23414
927203ad4b3a
tuned conjunction tactics: slightly smaller proof terms;
wenzelm
parents:
23356
diff
changeset
|
13 |
val PRECISE_CONJUNCTS: int -> tactic -> int -> tactic |
32365 | 14 |
val PARALLEL_CHOICE: tactic list -> tactic |
15 |
val PARALLEL_GOALS: tactic -> tactic |
|
17980 | 16 |
end; |
17 |
||
18 |
signature GOAL = |
|
19 |
sig |
|
20 |
include BASIC_GOAL |
|
21 |
val init: cterm -> thm |
|
18027
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
22 |
val protect: thm -> thm |
17980 | 23 |
val conclude: thm -> thm |
32197 | 24 |
val check_finished: Proof.context -> thm -> unit |
25 |
val finish: Proof.context -> thm -> thm |
|
21604
1af327306c8e
added norm/close_result (supercede local_standard etc.);
wenzelm
parents:
21579
diff
changeset
|
26 |
val norm_result: thm -> thm |
41677 | 27 |
val fork_name: string -> (unit -> 'a) -> 'a future |
37186
349e9223c685
explicit markup for forked goals, as indicated by Goal.fork;
wenzelm
parents:
36613
diff
changeset
|
28 |
val fork: (unit -> 'a) -> 'a future |
29448
34b9652b2f45
added Goal.future_enabled abstraction -- now also checks that this is already
wenzelm
parents:
29435
diff
changeset
|
29 |
val future_enabled: unit -> bool |
32061
11f8ee55662d
parallel_proofs: more fine-grained control with optional parallel checking of nested Isar proofs;
wenzelm
parents:
32058
diff
changeset
|
30 |
val local_future_enabled: unit -> bool |
28973
c549650d1442
future proofs: pass actual futures to facilitate composite computations;
wenzelm
parents:
28619
diff
changeset
|
31 |
val future_result: Proof.context -> thm future -> term -> thm |
23356 | 32 |
val prove_internal: cterm list -> cterm -> (thm list -> tactic) -> thm |
20290 | 33 |
val prove_multi: Proof.context -> string list -> term list -> term list -> |
34 |
({prems: thm list, context: Proof.context} -> tactic) -> thm list |
|
28446
a01de3b3fa2e
renamed promise to future, tuned related interfaces;
wenzelm
parents:
28363
diff
changeset
|
35 |
val prove_future: Proof.context -> string list -> term list -> term -> |
28340 | 36 |
({prems: thm list, context: Proof.context} -> tactic) -> thm |
20290 | 37 |
val prove: Proof.context -> string list -> term list -> term -> |
38 |
({prems: thm list, context: Proof.context} -> tactic) -> thm |
|
26713
1c6181def1d0
prove_global: Variable.set_body true, pass context;
wenzelm
parents:
26628
diff
changeset
|
39 |
val prove_global: theory -> string list -> term list -> term -> |
1c6181def1d0
prove_global: Variable.set_body true, pass context;
wenzelm
parents:
26628
diff
changeset
|
40 |
({prems: thm list, context: Proof.context} -> tactic) -> thm |
19184 | 41 |
val extract: int -> int -> thm -> thm Seq.seq |
42 |
val retrofit: int -> int -> thm -> thm -> thm Seq.seq |
|
23418 | 43 |
val conjunction_tac: int -> tactic |
23414
927203ad4b3a
tuned conjunction tactics: slightly smaller proof terms;
wenzelm
parents:
23356
diff
changeset
|
44 |
val precise_conjunction_tac: int -> int -> tactic |
23418 | 45 |
val recover_conjunction_tac: tactic |
21687 | 46 |
val norm_hhf_tac: int -> tactic |
47 |
val compose_hhf_tac: thm -> int -> tactic |
|
23237 | 48 |
val assume_rule_tac: Proof.context -> int -> tactic |
17980 | 49 |
end; |
50 |
||
51 |
structure Goal: GOAL = |
|
52 |
struct |
|
53 |
||
18027
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
54 |
(** goals **) |
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
55 |
|
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
56 |
(* |
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
57 |
-------- (init) |
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
58 |
C ==> #C |
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
59 |
*) |
20290 | 60 |
val init = |
22902
ac833b4bb7ee
moved some Drule operations to Thm (see more_thm.ML);
wenzelm
parents:
21687
diff
changeset
|
61 |
let val A = #1 (Thm.dest_implies (Thm.cprop_of Drule.protectI)) |
20290 | 62 |
in fn C => Thm.instantiate ([], [(A, C)]) Drule.protectI end; |
17980 | 63 |
|
64 |
(* |
|
18119 | 65 |
C |
66 |
--- (protect) |
|
67 |
#C |
|
17980 | 68 |
*) |
29345 | 69 |
fun protect th = Drule.comp_no_flatten (th, 0) 1 Drule.protectI; |
17980 | 70 |
|
71 |
(* |
|
18027
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
72 |
A ==> ... ==> #C |
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
73 |
---------------- (conclude) |
17980 | 74 |
A ==> ... ==> C |
75 |
*) |
|
29345 | 76 |
fun conclude th = Drule.comp_no_flatten (th, Thm.nprems_of th) 1 Drule.protectD; |
17980 | 77 |
|
78 |
(* |
|
18027
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
79 |
#C |
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
80 |
--- (finish) |
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
81 |
C |
17983 | 82 |
*) |
32197 | 83 |
fun check_finished ctxt th = |
17980 | 84 |
(case Thm.nprems_of th of |
32197 | 85 |
0 => () |
17980 | 86 |
| n => raise THM ("Proof failed.\n" ^ |
32197 | 87 |
Pretty.string_of (Pretty.chunks |
39125
f45d332a90e3
pretty_goals: turned some global references and function arguments into configuration options (goals_limit = 10, goals_total = true, show_main_goal = false) depending on the context;
wenzelm
parents:
38875
diff
changeset
|
88 |
(Goal_Display.pretty_goals |
f45d332a90e3
pretty_goals: turned some global references and function arguments into configuration options (goals_limit = 10, goals_total = true, show_main_goal = false) depending on the context;
wenzelm
parents:
38875
diff
changeset
|
89 |
(ctxt |
f45d332a90e3
pretty_goals: turned some global references and function arguments into configuration options (goals_limit = 10, goals_total = true, show_main_goal = false) depending on the context;
wenzelm
parents:
38875
diff
changeset
|
90 |
|> Config.put Goal_Display.goals_limit n |
f45d332a90e3
pretty_goals: turned some global references and function arguments into configuration options (goals_limit = 10, goals_total = true, show_main_goal = false) depending on the context;
wenzelm
parents:
38875
diff
changeset
|
91 |
|> Config.put Goal_Display.show_main_goal true) th)) ^ |
32197 | 92 |
"\n" ^ string_of_int n ^ " unsolved goal(s)!", 0, [th])); |
93 |
||
94 |
fun finish ctxt th = (check_finished ctxt th; conclude th); |
|
17980 | 95 |
|
96 |
||
18027
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
97 |
|
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
98 |
(** results **) |
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
99 |
|
28340 | 100 |
(* normal form *) |
101 |
||
21604
1af327306c8e
added norm/close_result (supercede local_standard etc.);
wenzelm
parents:
21579
diff
changeset
|
102 |
val norm_result = |
1af327306c8e
added norm/close_result (supercede local_standard etc.);
wenzelm
parents:
21579
diff
changeset
|
103 |
Drule.flexflex_unique |
41228
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents:
39125
diff
changeset
|
104 |
#> Raw_Simplifier.norm_hhf_protect |
21604
1af327306c8e
added norm/close_result (supercede local_standard etc.);
wenzelm
parents:
21579
diff
changeset
|
105 |
#> Thm.strip_shyps |
1af327306c8e
added norm/close_result (supercede local_standard etc.);
wenzelm
parents:
21579
diff
changeset
|
106 |
#> Drule.zero_var_indexes; |
1af327306c8e
added norm/close_result (supercede local_standard etc.);
wenzelm
parents:
21579
diff
changeset
|
107 |
|
1af327306c8e
added norm/close_result (supercede local_standard etc.);
wenzelm
parents:
21579
diff
changeset
|
108 |
|
41703
d27950860514
parallelization of nested Isar proofs is subject to Goal.parallel_proofs_threshold;
wenzelm
parents:
41683
diff
changeset
|
109 |
(* forked proofs *) |
d27950860514
parallelization of nested Isar proofs is subject to Goal.parallel_proofs_threshold;
wenzelm
parents:
41683
diff
changeset
|
110 |
|
d27950860514
parallelization of nested Isar proofs is subject to Goal.parallel_proofs_threshold;
wenzelm
parents:
41683
diff
changeset
|
111 |
val forked_proofs = Synchronized.var "forked_proofs" 0; |
d27950860514
parallelization of nested Isar proofs is subject to Goal.parallel_proofs_threshold;
wenzelm
parents:
41683
diff
changeset
|
112 |
|
d27950860514
parallelization of nested Isar proofs is subject to Goal.parallel_proofs_threshold;
wenzelm
parents:
41683
diff
changeset
|
113 |
fun forked i = |
d27950860514
parallelization of nested Isar proofs is subject to Goal.parallel_proofs_threshold;
wenzelm
parents:
41683
diff
changeset
|
114 |
Synchronized.change forked_proofs (fn m => |
d27950860514
parallelization of nested Isar proofs is subject to Goal.parallel_proofs_threshold;
wenzelm
parents:
41683
diff
changeset
|
115 |
let |
d27950860514
parallelization of nested Isar proofs is subject to Goal.parallel_proofs_threshold;
wenzelm
parents:
41683
diff
changeset
|
116 |
val n = m + i; |
d27950860514
parallelization of nested Isar proofs is subject to Goal.parallel_proofs_threshold;
wenzelm
parents:
41683
diff
changeset
|
117 |
val _ = |
41776 | 118 |
Multithreading.tracing 2 (fn () => |
41703
d27950860514
parallelization of nested Isar proofs is subject to Goal.parallel_proofs_threshold;
wenzelm
parents:
41683
diff
changeset
|
119 |
("PROOFS " ^ Time.toString (Time.now ()) ^ ": " ^ string_of_int n)); |
d27950860514
parallelization of nested Isar proofs is subject to Goal.parallel_proofs_threshold;
wenzelm
parents:
41683
diff
changeset
|
120 |
in n end); |
37186
349e9223c685
explicit markup for forked goals, as indicated by Goal.fork;
wenzelm
parents:
36613
diff
changeset
|
121 |
|
41677 | 122 |
fun fork_name name e = |
41703
d27950860514
parallelization of nested Isar proofs is subject to Goal.parallel_proofs_threshold;
wenzelm
parents:
41683
diff
changeset
|
123 |
uninterruptible (fn _ => fn () => |
d27950860514
parallelization of nested Isar proofs is subject to Goal.parallel_proofs_threshold;
wenzelm
parents:
41683
diff
changeset
|
124 |
let |
d27950860514
parallelization of nested Isar proofs is subject to Goal.parallel_proofs_threshold;
wenzelm
parents:
41683
diff
changeset
|
125 |
val _ = forked 1; |
d27950860514
parallelization of nested Isar proofs is subject to Goal.parallel_proofs_threshold;
wenzelm
parents:
41683
diff
changeset
|
126 |
val future = |
44302 | 127 |
(singleton o Future.forks) |
45458
5b5d3ee2285b
depth-first proof forking for improved locality (wrt. cancellation and overall memory usage);
wenzelm
parents:
45344
diff
changeset
|
128 |
{name = name, group = NONE, deps = [], pri = 0, interrupts = false} |
41703
d27950860514
parallelization of nested Isar proofs is subject to Goal.parallel_proofs_threshold;
wenzelm
parents:
41683
diff
changeset
|
129 |
(fn () => |
d27950860514
parallelization of nested Isar proofs is subject to Goal.parallel_proofs_threshold;
wenzelm
parents:
41683
diff
changeset
|
130 |
Exn.release |
44114 | 131 |
(Exn.capture (Future.status o Future.interruptible_task) e before forked ~1)); |
41703
d27950860514
parallelization of nested Isar proofs is subject to Goal.parallel_proofs_threshold;
wenzelm
parents:
41683
diff
changeset
|
132 |
in future end) (); |
29448
34b9652b2f45
added Goal.future_enabled abstraction -- now also checks that this is already
wenzelm
parents:
29435
diff
changeset
|
133 |
|
41677 | 134 |
fun fork e = fork_name "Goal.fork" e; |
135 |
||
136 |
||
41703
d27950860514
parallelization of nested Isar proofs is subject to Goal.parallel_proofs_threshold;
wenzelm
parents:
41683
diff
changeset
|
137 |
(* scheduling parameters *) |
d27950860514
parallelization of nested Isar proofs is subject to Goal.parallel_proofs_threshold;
wenzelm
parents:
41683
diff
changeset
|
138 |
|
32738 | 139 |
val parallel_proofs = Unsynchronized.ref 1; |
41819
2d84831dc1a0
scale parallel_proofs_threshold with max_threads_value to improve saturation of cores;
wenzelm
parents:
41776
diff
changeset
|
140 |
val parallel_proofs_threshold = Unsynchronized.ref 50; |
29448
34b9652b2f45
added Goal.future_enabled abstraction -- now also checks that this is already
wenzelm
parents:
29435
diff
changeset
|
141 |
|
34b9652b2f45
added Goal.future_enabled abstraction -- now also checks that this is already
wenzelm
parents:
29435
diff
changeset
|
142 |
fun future_enabled () = |
41703
d27950860514
parallelization of nested Isar proofs is subject to Goal.parallel_proofs_threshold;
wenzelm
parents:
41683
diff
changeset
|
143 |
Multithreading.enabled () andalso ! parallel_proofs >= 1 andalso |
d27950860514
parallelization of nested Isar proofs is subject to Goal.parallel_proofs_threshold;
wenzelm
parents:
41683
diff
changeset
|
144 |
is_some (Future.worker_task ()); |
32061
11f8ee55662d
parallel_proofs: more fine-grained control with optional parallel checking of nested Isar proofs;
wenzelm
parents:
32058
diff
changeset
|
145 |
|
41703
d27950860514
parallelization of nested Isar proofs is subject to Goal.parallel_proofs_threshold;
wenzelm
parents:
41683
diff
changeset
|
146 |
fun local_future_enabled () = |
d27950860514
parallelization of nested Isar proofs is subject to Goal.parallel_proofs_threshold;
wenzelm
parents:
41683
diff
changeset
|
147 |
future_enabled () andalso ! parallel_proofs >= 2 andalso |
41819
2d84831dc1a0
scale parallel_proofs_threshold with max_threads_value to improve saturation of cores;
wenzelm
parents:
41776
diff
changeset
|
148 |
Synchronized.value forked_proofs < |
2d84831dc1a0
scale parallel_proofs_threshold with max_threads_value to improve saturation of cores;
wenzelm
parents:
41776
diff
changeset
|
149 |
! parallel_proofs_threshold * Multithreading.max_threads_value (); |
29448
34b9652b2f45
added Goal.future_enabled abstraction -- now also checks that this is already
wenzelm
parents:
29435
diff
changeset
|
150 |
|
34b9652b2f45
added Goal.future_enabled abstraction -- now also checks that this is already
wenzelm
parents:
29435
diff
changeset
|
151 |
|
28446
a01de3b3fa2e
renamed promise to future, tuned related interfaces;
wenzelm
parents:
28363
diff
changeset
|
152 |
(* future_result *) |
28340 | 153 |
|
28446
a01de3b3fa2e
renamed promise to future, tuned related interfaces;
wenzelm
parents:
28363
diff
changeset
|
154 |
fun future_result ctxt result prop = |
28340 | 155 |
let |
42360 | 156 |
val thy = Proof_Context.theory_of ctxt; |
28340 | 157 |
val _ = Context.reject_draft thy; |
158 |
val cert = Thm.cterm_of thy; |
|
159 |
val certT = Thm.ctyp_of thy; |
|
160 |
||
30473
e0b66c11e7e4
Assumption.all_prems_of, Assumption.all_assms_of;
wenzelm
parents:
29448
diff
changeset
|
161 |
val assms = Assumption.all_assms_of ctxt; |
28340 | 162 |
val As = map Thm.term_of assms; |
163 |
||
164 |
val xs = map Free (fold Term.add_frees (prop :: As) []); |
|
165 |
val fixes = map cert xs; |
|
166 |
||
167 |
val tfrees = fold Term.add_tfrees (prop :: As) []; |
|
168 |
val instT = map (fn (a, S) => (certT (TVar ((a, 0), S)), certT (TFree (a, S)))) tfrees; |
|
169 |
||
170 |
val global_prop = |
|
45344
e209da839ff4
added Logic.varify_types_global/unvarify_types_global, which avoids somewhat expensive Term.map_types;
wenzelm
parents:
44302
diff
changeset
|
171 |
cert (Logic.varify_types_global (fold_rev Logic.all xs (Logic.list_implies (As, prop)))) |
32056
f4b74cbecdaf
future_result: explicitly impose Variable.sorts_of again;
wenzelm
parents:
30473
diff
changeset
|
172 |
|> Thm.weaken_sorts (Variable.sorts_of ctxt); |
28973
c549650d1442
future proofs: pass actual futures to facilitate composite computations;
wenzelm
parents:
28619
diff
changeset
|
173 |
val global_result = result |> Future.map |
33768
bba9eac8aa25
future_result: purge flexflex pairs, which should normally be trivial anyway -- prevent Thm.future_result from complaining about tpairs;
wenzelm
parents:
32788
diff
changeset
|
174 |
(Drule.flexflex_unique #> |
bba9eac8aa25
future_result: purge flexflex pairs, which should normally be trivial anyway -- prevent Thm.future_result from complaining about tpairs;
wenzelm
parents:
32788
diff
changeset
|
175 |
Thm.adjust_maxidx_thm ~1 #> |
28973
c549650d1442
future proofs: pass actual futures to facilitate composite computations;
wenzelm
parents:
28619
diff
changeset
|
176 |
Drule.implies_intr_list assms #> |
c549650d1442
future proofs: pass actual futures to facilitate composite computations;
wenzelm
parents:
28619
diff
changeset
|
177 |
Drule.forall_intr_list fixes #> |
36613
f3157c288aca
simplified primitive Thm.future: more direct theory check, do not hardwire strip_shyps here;
wenzelm
parents:
36610
diff
changeset
|
178 |
Thm.generalize (map #1 tfrees, []) 0 #> |
f3157c288aca
simplified primitive Thm.future: more direct theory check, do not hardwire strip_shyps here;
wenzelm
parents:
36610
diff
changeset
|
179 |
Thm.strip_shyps); |
28340 | 180 |
val local_result = |
32056
f4b74cbecdaf
future_result: explicitly impose Variable.sorts_of again;
wenzelm
parents:
30473
diff
changeset
|
181 |
Thm.future global_result global_prop |
28340 | 182 |
|> Thm.instantiate (instT, []) |
183 |
|> Drule.forall_elim_list fixes |
|
184 |
|> fold (Thm.elim_implies o Thm.assume) assms; |
|
185 |
in local_result end; |
|
186 |
||
187 |
||
18027
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
188 |
|
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
189 |
(** tactical theorem proving **) |
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
190 |
|
23356 | 191 |
(* prove_internal -- minimal checks, no normalization of result! *) |
20250
c3f209752749
prove: proper assumption context, more tactic arguments;
wenzelm
parents:
20228
diff
changeset
|
192 |
|
23356 | 193 |
fun prove_internal casms cprop tac = |
20250
c3f209752749
prove: proper assumption context, more tactic arguments;
wenzelm
parents:
20228
diff
changeset
|
194 |
(case SINGLE (tac (map Assumption.assume casms)) (init cprop) of |
32197 | 195 |
SOME th => Drule.implies_intr_list casms |
196 |
(finish (Syntax.init_pretty_global (Thm.theory_of_thm th)) th) |
|
38875
c7a66b584147
tuned messages: discontinued spurious full-stops (messages are occasionally composed unexpectedly);
wenzelm
parents:
38236
diff
changeset
|
197 |
| NONE => error "Tactic failed"); |
20250
c3f209752749
prove: proper assumption context, more tactic arguments;
wenzelm
parents:
20228
diff
changeset
|
198 |
|
c3f209752749
prove: proper assumption context, more tactic arguments;
wenzelm
parents:
20228
diff
changeset
|
199 |
|
28340 | 200 |
(* prove_common etc. *) |
17986 | 201 |
|
28340 | 202 |
fun prove_common immediate ctxt xs asms props tac = |
17980 | 203 |
let |
42360 | 204 |
val thy = Proof_Context.theory_of ctxt; |
26939
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents:
26713
diff
changeset
|
205 |
val string_of_term = Syntax.string_of_term ctxt; |
20056 | 206 |
|
28355 | 207 |
val pos = Position.thread_data (); |
20250
c3f209752749
prove: proper assumption context, more tactic arguments;
wenzelm
parents:
20228
diff
changeset
|
208 |
fun err msg = cat_error msg |
c3f209752749
prove: proper assumption context, more tactic arguments;
wenzelm
parents:
20228
diff
changeset
|
209 |
("The error(s) above occurred for the goal statement:\n" ^ |
28355 | 210 |
string_of_term (Logic.list_implies (asms, Logic.mk_conjunction_list props)) ^ |
211 |
(case Position.str_of pos of "" => "" | s => "\n" ^ s)); |
|
17980 | 212 |
|
20250
c3f209752749
prove: proper assumption context, more tactic arguments;
wenzelm
parents:
20228
diff
changeset
|
213 |
fun cert_safe t = Thm.cterm_of thy (Envir.beta_norm (Term.no_dummy_patterns t)) |
17980 | 214 |
handle TERM (msg, _) => err msg | TYPE (msg, _, _) => err msg; |
20250
c3f209752749
prove: proper assumption context, more tactic arguments;
wenzelm
parents:
20228
diff
changeset
|
215 |
val casms = map cert_safe asms; |
c3f209752749
prove: proper assumption context, more tactic arguments;
wenzelm
parents:
20228
diff
changeset
|
216 |
val cprops = map cert_safe props; |
17980 | 217 |
|
20250
c3f209752749
prove: proper assumption context, more tactic arguments;
wenzelm
parents:
20228
diff
changeset
|
218 |
val (prems, ctxt') = ctxt |
c3f209752749
prove: proper assumption context, more tactic arguments;
wenzelm
parents:
20228
diff
changeset
|
219 |
|> Variable.add_fixes_direct xs |
27218
4548c83cd508
prove: full Variable.declare_term, including constraints;
wenzelm
parents:
26939
diff
changeset
|
220 |
|> fold Variable.declare_term (asms @ props) |
26713
1c6181def1d0
prove_global: Variable.set_body true, pass context;
wenzelm
parents:
26628
diff
changeset
|
221 |
|> Assumption.add_assumes casms |
1c6181def1d0
prove_global: Variable.set_body true, pass context;
wenzelm
parents:
26628
diff
changeset
|
222 |
||> Variable.set_body true; |
28619
89f9dd800a22
prove_common: include all sorts from context into statement, check shyps of result;
wenzelm
parents:
28446
diff
changeset
|
223 |
val sorts = Variable.sorts_of ctxt'; |
17980 | 224 |
|
28619
89f9dd800a22
prove_common: include all sorts from context into statement, check shyps of result;
wenzelm
parents:
28446
diff
changeset
|
225 |
val stmt = Thm.weaken_sorts sorts (Conjunction.mk_conjunction_balanced cprops); |
28340 | 226 |
|
227 |
fun result () = |
|
228 |
(case SINGLE (tac {prems = prems, context = ctxt'}) (init stmt) of |
|
38875
c7a66b584147
tuned messages: discontinued spurious full-stops (messages are occasionally composed unexpectedly);
wenzelm
parents:
38236
diff
changeset
|
229 |
NONE => err "Tactic failed" |
28340 | 230 |
| SOME st => |
32197 | 231 |
let val res = finish ctxt' st handle THM (msg, _, _) => err msg in |
28619
89f9dd800a22
prove_common: include all sorts from context into statement, check shyps of result;
wenzelm
parents:
28446
diff
changeset
|
232 |
if Unify.matches_list thy [Thm.term_of stmt] [Thm.prop_of res] |
89f9dd800a22
prove_common: include all sorts from context into statement, check shyps of result;
wenzelm
parents:
28446
diff
changeset
|
233 |
then Thm.check_shyps sorts res |
28340 | 234 |
else err ("Proved a different theorem: " ^ string_of_term (Thm.prop_of res)) |
235 |
end); |
|
236 |
val res = |
|
29448
34b9652b2f45
added Goal.future_enabled abstraction -- now also checks that this is already
wenzelm
parents:
29435
diff
changeset
|
237 |
if immediate orelse #maxidx (Thm.rep_cterm stmt) >= 0 orelse not (future_enabled ()) |
29088
95a239a5e055
future proofs: more robust check via Future.enabled;
wenzelm
parents:
28973
diff
changeset
|
238 |
then result () |
37186
349e9223c685
explicit markup for forked goals, as indicated by Goal.fork;
wenzelm
parents:
36613
diff
changeset
|
239 |
else future_result ctxt' (fork result) (Thm.term_of stmt); |
17980 | 240 |
in |
28340 | 241 |
Conjunction.elim_balanced (length props) res |
20290 | 242 |
|> map (Assumption.export false ctxt' ctxt) |
20056 | 243 |
|> Variable.export ctxt' ctxt |
20250
c3f209752749
prove: proper assumption context, more tactic arguments;
wenzelm
parents:
20228
diff
changeset
|
244 |
|> map Drule.zero_var_indexes |
17980 | 245 |
end; |
246 |
||
28341 | 247 |
val prove_multi = prove_common true; |
17980 | 248 |
|
28446
a01de3b3fa2e
renamed promise to future, tuned related interfaces;
wenzelm
parents:
28363
diff
changeset
|
249 |
fun prove_future ctxt xs asms prop tac = hd (prove_common false ctxt xs asms [prop] tac); |
28340 | 250 |
fun prove ctxt xs asms prop tac = hd (prove_common true ctxt xs asms [prop] tac); |
20056 | 251 |
|
252 |
fun prove_global thy xs asms prop tac = |
|
42360 | 253 |
Drule.export_without_context (prove (Proof_Context.init_global thy) xs asms prop tac); |
18027
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
254 |
|
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents:
17986
diff
changeset
|
255 |
|
17980 | 256 |
|
21687 | 257 |
(** goal structure **) |
258 |
||
259 |
(* nested goals *) |
|
18207 | 260 |
|
19184 | 261 |
fun extract i n st = |
262 |
(if i < 1 orelse n < 1 orelse i + n - 1 > Thm.nprems_of st then Seq.empty |
|
263 |
else if n = 1 then Seq.single (Thm.cprem_of st i) |
|
23418 | 264 |
else |
265 |
Seq.single (Conjunction.mk_conjunction_balanced (map (Thm.cprem_of st) (i upto i + n - 1)))) |
|
20260 | 266 |
|> Seq.map (Thm.adjust_maxidx_cterm ~1 #> init); |
17980 | 267 |
|
19221 | 268 |
fun retrofit i n st' st = |
269 |
(if n = 1 then st |
|
23418 | 270 |
else st |> Drule.with_subgoal i (Conjunction.uncurry_balanced n)) |
19221 | 271 |
|> Thm.compose_no_flatten false (conclude st', Thm.nprems_of st') i; |
18207 | 272 |
|
17980 | 273 |
fun SELECT_GOAL tac i st = |
19191 | 274 |
if Thm.nprems_of st = 1 andalso i = 1 then tac st |
19184 | 275 |
else Seq.lifts (retrofit i 1) (Seq.maps tac (extract i 1 st)) st; |
17980 | 276 |
|
21687 | 277 |
|
278 |
(* multiple goals *) |
|
279 |
||
23418 | 280 |
fun precise_conjunction_tac 0 i = eq_assume_tac i |
281 |
| precise_conjunction_tac 1 i = SUBGOAL (K all_tac) i |
|
282 |
| precise_conjunction_tac n i = PRIMITIVE (Drule.with_subgoal i (Conjunction.curry_balanced n)); |
|
23414
927203ad4b3a
tuned conjunction tactics: slightly smaller proof terms;
wenzelm
parents:
23356
diff
changeset
|
283 |
|
23418 | 284 |
val adhoc_conjunction_tac = REPEAT_ALL_NEW |
285 |
(SUBGOAL (fn (goal, i) => |
|
286 |
if can Logic.dest_conjunction goal then rtac Conjunction.conjunctionI i |
|
287 |
else no_tac)); |
|
21687 | 288 |
|
23418 | 289 |
val conjunction_tac = SUBGOAL (fn (goal, i) => |
290 |
precise_conjunction_tac (length (Logic.dest_conjunctions goal)) i ORELSE |
|
291 |
TRY (adhoc_conjunction_tac i)); |
|
21687 | 292 |
|
23418 | 293 |
val recover_conjunction_tac = PRIMITIVE (fn th => |
294 |
Conjunction.uncurry_balanced (Thm.nprems_of th) th); |
|
23414
927203ad4b3a
tuned conjunction tactics: slightly smaller proof terms;
wenzelm
parents:
23356
diff
changeset
|
295 |
|
927203ad4b3a
tuned conjunction tactics: slightly smaller proof terms;
wenzelm
parents:
23356
diff
changeset
|
296 |
fun PRECISE_CONJUNCTS n tac = |
927203ad4b3a
tuned conjunction tactics: slightly smaller proof terms;
wenzelm
parents:
23356
diff
changeset
|
297 |
SELECT_GOAL (precise_conjunction_tac n 1 |
927203ad4b3a
tuned conjunction tactics: slightly smaller proof terms;
wenzelm
parents:
23356
diff
changeset
|
298 |
THEN tac |
23418 | 299 |
THEN recover_conjunction_tac); |
23414
927203ad4b3a
tuned conjunction tactics: slightly smaller proof terms;
wenzelm
parents:
23356
diff
changeset
|
300 |
|
21687 | 301 |
fun CONJUNCTS tac = |
302 |
SELECT_GOAL (conjunction_tac 1 |
|
303 |
THEN tac |
|
23418 | 304 |
THEN recover_conjunction_tac); |
21687 | 305 |
|
306 |
||
307 |
(* hhf normal form *) |
|
308 |
||
309 |
val norm_hhf_tac = |
|
310 |
rtac Drule.asm_rl (*cheap approximation -- thanks to builtin Logic.flatten_params*) |
|
311 |
THEN' SUBGOAL (fn (t, i) => |
|
312 |
if Drule.is_norm_hhf t then all_tac |
|
41228
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents:
39125
diff
changeset
|
313 |
else rewrite_goal_tac Drule.norm_hhf_eqs i); |
21687 | 314 |
|
25301 | 315 |
fun compose_hhf_tac th i st = |
316 |
PRIMSEQ (Thm.bicompose false (false, Drule.lift_all (Thm.cprem_of st i) th, 0) i) st; |
|
21687 | 317 |
|
23237 | 318 |
|
319 |
(* non-atomic goal assumptions *) |
|
320 |
||
23356 | 321 |
fun non_atomic (Const ("==>", _) $ _ $ _) = true |
322 |
| non_atomic (Const ("all", _) $ _) = true |
|
323 |
| non_atomic _ = false; |
|
324 |
||
23237 | 325 |
fun assume_rule_tac ctxt = norm_hhf_tac THEN' CSUBGOAL (fn (goal, i) => |
326 |
let |
|
42495
1af81b70cf09
clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
wenzelm
parents:
42371
diff
changeset
|
327 |
val ((_, goal'), ctxt') = Variable.focus_cterm goal ctxt; |
23237 | 328 |
val goal'' = Drule.cterm_rule (singleton (Variable.export ctxt' ctxt)) goal'; |
23356 | 329 |
val Rs = filter (non_atomic o Thm.term_of) (Drule.strip_imp_prems goal''); |
23237 | 330 |
val tacs = Rs |> map (fn R => |
41228
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
wenzelm
parents:
39125
diff
changeset
|
331 |
Tactic.etac (Raw_Simplifier.norm_hhf (Thm.trivial R)) THEN_ALL_NEW assume_tac); |
23237 | 332 |
in fold_rev (curry op APPEND') tacs (K no_tac) i end); |
333 |
||
32365 | 334 |
|
335 |
(* parallel tacticals *) |
|
336 |
||
337 |
(*parallel choice of single results*) |
|
338 |
fun PARALLEL_CHOICE tacs st = |
|
339 |
(case Par_List.get_some (fn tac => SINGLE tac st) tacs of |
|
340 |
NONE => Seq.empty |
|
341 |
| SOME st' => Seq.single st'); |
|
342 |
||
343 |
(*parallel refinement of non-schematic goal by single results*) |
|
32788
a65deb8f9434
PARALLEL_GOALS: proper scope for exception FAILED, with dummy argument to prevent its interpretation as variable;
wenzelm
parents:
32738
diff
changeset
|
344 |
exception FAILED of unit; |
42370 | 345 |
fun PARALLEL_GOALS tac = |
346 |
Thm.adjust_maxidx_thm ~1 #> |
|
347 |
(fn st => |
|
42371
5900f06b4198
enable PARALLEL_GOALS more liberally, unlike forked proofs (cf. 34b9652b2f45);
wenzelm
parents:
42370
diff
changeset
|
348 |
if not (Multithreading.enabled ()) orelse Thm.maxidx_of st >= 0 orelse Thm.nprems_of st <= 1 |
42370 | 349 |
then DETERM tac st |
350 |
else |
|
351 |
let |
|
352 |
fun try_tac g = |
|
353 |
(case SINGLE tac g of |
|
354 |
NONE => raise FAILED () |
|
355 |
| SOME g' => g'); |
|
32365 | 356 |
|
42370 | 357 |
val goals = Drule.strip_imp_prems (Thm.cprop_of st); |
358 |
val results = Par_List.map (try_tac o init) goals; |
|
359 |
in ALLGOALS (fn i => retrofit i 1 (nth results (i - 1))) st end |
|
360 |
handle FAILED () => Seq.empty); |
|
32365 | 361 |
|
18207 | 362 |
end; |
363 |
||
32365 | 364 |
structure Basic_Goal: BASIC_GOAL = Goal; |
365 |
open Basic_Goal; |