| author | wenzelm | 
| Sat, 15 Jan 2011 14:19:37 +0100 | |
| changeset 41573 | 7a3181986004 | 
| parent 41228 | e1fce873b814 | 
| child 41673 | 1c191a39549f | 
| 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  | 
| 17980 | 10  | 
val SELECT_GOAL: tactic -> int -> tactic  | 
| 23418 | 11  | 
val CONJUNCTS: tactic -> int -> tactic  | 
| 
23414
 
927203ad4b3a
tuned conjunction tactics: slightly smaller proof terms;
 
wenzelm 
parents: 
23356 
diff
changeset
 | 
12  | 
val PRECISE_CONJUNCTS: int -> tactic -> int -> tactic  | 
| 32365 | 13  | 
val PARALLEL_CHOICE: tactic list -> tactic  | 
14  | 
val PARALLEL_GOALS: tactic -> tactic  | 
|
| 17980 | 15  | 
end;  | 
16  | 
||
17  | 
signature GOAL =  | 
|
18  | 
sig  | 
|
19  | 
include BASIC_GOAL  | 
|
20  | 
val init: cterm -> thm  | 
|
| 
18027
 
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
 
wenzelm 
parents: 
17986 
diff
changeset
 | 
21  | 
val protect: thm -> thm  | 
| 17980 | 22  | 
val conclude: thm -> thm  | 
| 32197 | 23  | 
val check_finished: Proof.context -> thm -> unit  | 
24  | 
val finish: Proof.context -> thm -> thm  | 
|
| 
21604
 
1af327306c8e
added norm/close_result (supercede local_standard etc.);
 
wenzelm 
parents: 
21579 
diff
changeset
 | 
25  | 
val norm_result: thm -> thm  | 
| 
37186
 
349e9223c685
explicit markup for forked goals, as indicated by Goal.fork;
 
wenzelm 
parents: 
36613 
diff
changeset
 | 
26  | 
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
 | 
27  | 
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
 | 
28  | 
val local_future_enabled: unit -> bool  | 
| 
28973
 
c549650d1442
future proofs: pass actual futures to facilitate composite computations;
 
wenzelm 
parents: 
28619 
diff
changeset
 | 
29  | 
val future_result: Proof.context -> thm future -> term -> thm  | 
| 23356 | 30  | 
val prove_internal: cterm list -> cterm -> (thm list -> tactic) -> thm  | 
| 20290 | 31  | 
val prove_multi: Proof.context -> string list -> term list -> term list ->  | 
32  | 
    ({prems: thm list, context: Proof.context} -> tactic) -> thm list
 | 
|
| 
28446
 
a01de3b3fa2e
renamed promise to future, tuned related interfaces;
 
wenzelm 
parents: 
28363 
diff
changeset
 | 
33  | 
val prove_future: Proof.context -> string list -> term list -> term ->  | 
| 28340 | 34  | 
    ({prems: thm list, context: Proof.context} -> tactic) -> thm
 | 
| 20290 | 35  | 
val prove: Proof.context -> string list -> term list -> term ->  | 
36  | 
    ({prems: thm list, context: Proof.context} -> tactic) -> thm
 | 
|
| 
26713
 
1c6181def1d0
prove_global: Variable.set_body true, pass context;
 
wenzelm 
parents: 
26628 
diff
changeset
 | 
37  | 
val prove_global: theory -> string list -> term list -> term ->  | 
| 
 
1c6181def1d0
prove_global: Variable.set_body true, pass context;
 
wenzelm 
parents: 
26628 
diff
changeset
 | 
38  | 
    ({prems: thm list, context: Proof.context} -> tactic) -> thm
 | 
| 19184 | 39  | 
val extract: int -> int -> thm -> thm Seq.seq  | 
40  | 
val retrofit: int -> int -> thm -> thm -> thm Seq.seq  | 
|
| 23418 | 41  | 
val conjunction_tac: int -> tactic  | 
| 
23414
 
927203ad4b3a
tuned conjunction tactics: slightly smaller proof terms;
 
wenzelm 
parents: 
23356 
diff
changeset
 | 
42  | 
val precise_conjunction_tac: int -> int -> tactic  | 
| 23418 | 43  | 
val recover_conjunction_tac: tactic  | 
| 21687 | 44  | 
val norm_hhf_tac: int -> tactic  | 
45  | 
val compose_hhf_tac: thm -> int -> tactic  | 
|
| 23237 | 46  | 
val assume_rule_tac: Proof.context -> int -> tactic  | 
| 17980 | 47  | 
end;  | 
48  | 
||
49  | 
structure Goal: GOAL =  | 
|
50  | 
struct  | 
|
51  | 
||
| 
18027
 
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
 
wenzelm 
parents: 
17986 
diff
changeset
 | 
52  | 
(** goals **)  | 
| 
 
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
 
wenzelm 
parents: 
17986 
diff
changeset
 | 
53  | 
|
| 
 
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
 
wenzelm 
parents: 
17986 
diff
changeset
 | 
54  | 
(*  | 
| 
 
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
 
wenzelm 
parents: 
17986 
diff
changeset
 | 
55  | 
-------- (init)  | 
| 
 
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
 
wenzelm 
parents: 
17986 
diff
changeset
 | 
56  | 
C ==> #C  | 
| 
 
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
 
wenzelm 
parents: 
17986 
diff
changeset
 | 
57  | 
*)  | 
| 20290 | 58  | 
val init =  | 
| 
22902
 
ac833b4bb7ee
moved some Drule operations to Thm (see more_thm.ML);
 
wenzelm 
parents: 
21687 
diff
changeset
 | 
59  | 
let val A = #1 (Thm.dest_implies (Thm.cprop_of Drule.protectI))  | 
| 20290 | 60  | 
in fn C => Thm.instantiate ([], [(A, C)]) Drule.protectI end;  | 
| 17980 | 61  | 
|
62  | 
(*  | 
|
| 18119 | 63  | 
C  | 
64  | 
--- (protect)  | 
|
65  | 
#C  | 
|
| 17980 | 66  | 
*)  | 
| 29345 | 67  | 
fun protect th = Drule.comp_no_flatten (th, 0) 1 Drule.protectI;  | 
| 17980 | 68  | 
|
69  | 
(*  | 
|
| 
18027
 
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
 
wenzelm 
parents: 
17986 
diff
changeset
 | 
70  | 
A ==> ... ==> #C  | 
| 
 
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
 
wenzelm 
parents: 
17986 
diff
changeset
 | 
71  | 
---------------- (conclude)  | 
| 17980 | 72  | 
A ==> ... ==> C  | 
73  | 
*)  | 
|
| 29345 | 74  | 
fun conclude th = Drule.comp_no_flatten (th, Thm.nprems_of th) 1 Drule.protectD;  | 
| 17980 | 75  | 
|
76  | 
(*  | 
|
| 
18027
 
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
 
wenzelm 
parents: 
17986 
diff
changeset
 | 
77  | 
#C  | 
| 
 
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
 
wenzelm 
parents: 
17986 
diff
changeset
 | 
78  | 
--- (finish)  | 
| 
 
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
 
wenzelm 
parents: 
17986 
diff
changeset
 | 
79  | 
C  | 
| 17983 | 80  | 
*)  | 
| 32197 | 81  | 
fun check_finished ctxt th =  | 
| 17980 | 82  | 
(case Thm.nprems_of th of  | 
| 32197 | 83  | 
0 => ()  | 
| 17980 | 84  | 
  | n => raise THM ("Proof failed.\n" ^
 | 
| 32197 | 85  | 
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
 | 
86  | 
(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
 | 
87  | 
(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
 | 
88  | 
|> 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
 | 
89  | 
|> Config.put Goal_Display.show_main_goal true) th)) ^  | 
| 32197 | 90  | 
"\n" ^ string_of_int n ^ " unsolved goal(s)!", 0, [th]));  | 
91  | 
||
92  | 
fun finish ctxt th = (check_finished ctxt th; conclude th);  | 
|
| 17980 | 93  | 
|
94  | 
||
| 
18027
 
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
 
wenzelm 
parents: 
17986 
diff
changeset
 | 
95  | 
|
| 
 
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
 
wenzelm 
parents: 
17986 
diff
changeset
 | 
96  | 
(** results **)  | 
| 
 
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
 
wenzelm 
parents: 
17986 
diff
changeset
 | 
97  | 
|
| 28340 | 98  | 
(* normal form *)  | 
99  | 
||
| 
21604
 
1af327306c8e
added norm/close_result (supercede local_standard etc.);
 
wenzelm 
parents: 
21579 
diff
changeset
 | 
100  | 
val norm_result =  | 
| 
 
1af327306c8e
added norm/close_result (supercede local_standard etc.);
 
wenzelm 
parents: 
21579 
diff
changeset
 | 
101  | 
Drule.flexflex_unique  | 
| 
41228
 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 
wenzelm 
parents: 
39125 
diff
changeset
 | 
102  | 
#> Raw_Simplifier.norm_hhf_protect  | 
| 
21604
 
1af327306c8e
added norm/close_result (supercede local_standard etc.);
 
wenzelm 
parents: 
21579 
diff
changeset
 | 
103  | 
#> Thm.strip_shyps  | 
| 
 
1af327306c8e
added norm/close_result (supercede local_standard etc.);
 
wenzelm 
parents: 
21579 
diff
changeset
 | 
104  | 
#> Drule.zero_var_indexes;  | 
| 
 
1af327306c8e
added norm/close_result (supercede local_standard etc.);
 
wenzelm 
parents: 
21579 
diff
changeset
 | 
105  | 
|
| 
 
1af327306c8e
added norm/close_result (supercede local_standard etc.);
 
wenzelm 
parents: 
21579 
diff
changeset
 | 
106  | 
|
| 
37186
 
349e9223c685
explicit markup for forked goals, as indicated by Goal.fork;
 
wenzelm 
parents: 
36613 
diff
changeset
 | 
107  | 
(* parallel proofs *)  | 
| 
 
349e9223c685
explicit markup for forked goals, as indicated by Goal.fork;
 
wenzelm 
parents: 
36613 
diff
changeset
 | 
108  | 
|
| 
38236
 
d8c7be27e01d
explicitly distinguish Output.status (essential feedback) vs. Output.report (useful markup);
 
wenzelm 
parents: 
38218 
diff
changeset
 | 
109  | 
fun fork e = Future.fork_pri ~1 (fn () => Future.status e);  | 
| 
29448
 
34b9652b2f45
added Goal.future_enabled abstraction -- now also checks that this is already
 
wenzelm 
parents: 
29435 
diff
changeset
 | 
110  | 
|
| 32738 | 111  | 
val parallel_proofs = Unsynchronized.ref 1;  | 
| 
29448
 
34b9652b2f45
added Goal.future_enabled abstraction -- now also checks that this is already
 
wenzelm 
parents: 
29435 
diff
changeset
 | 
112  | 
|
| 
 
34b9652b2f45
added Goal.future_enabled abstraction -- now also checks that this is already
 
wenzelm 
parents: 
29435 
diff
changeset
 | 
113  | 
fun future_enabled () =  | 
| 
32255
 
d302f1c9e356
eliminated separate Future.enabled -- let Future.join fail explicitly in critical section, instead of entering sequential mode silently;
 
wenzelm 
parents: 
32197 
diff
changeset
 | 
114  | 
Multithreading.enabled () andalso Future.is_worker () andalso ! parallel_proofs >= 1;  | 
| 
32061
 
11f8ee55662d
parallel_proofs: more fine-grained control with optional parallel checking of nested Isar proofs;
 
wenzelm 
parents: 
32058 
diff
changeset
 | 
115  | 
|
| 
 
11f8ee55662d
parallel_proofs: more fine-grained control with optional parallel checking of nested Isar proofs;
 
wenzelm 
parents: 
32058 
diff
changeset
 | 
116  | 
fun local_future_enabled () = future_enabled () andalso ! parallel_proofs >= 2;  | 
| 
29448
 
34b9652b2f45
added Goal.future_enabled abstraction -- now also checks that this is already
 
wenzelm 
parents: 
29435 
diff
changeset
 | 
117  | 
|
| 
 
34b9652b2f45
added Goal.future_enabled abstraction -- now also checks that this is already
 
wenzelm 
parents: 
29435 
diff
changeset
 | 
118  | 
|
| 
28446
 
a01de3b3fa2e
renamed promise to future, tuned related interfaces;
 
wenzelm 
parents: 
28363 
diff
changeset
 | 
119  | 
(* future_result *)  | 
| 28340 | 120  | 
|
| 
28446
 
a01de3b3fa2e
renamed promise to future, tuned related interfaces;
 
wenzelm 
parents: 
28363 
diff
changeset
 | 
121  | 
fun future_result ctxt result prop =  | 
| 28340 | 122  | 
let  | 
123  | 
val thy = ProofContext.theory_of ctxt;  | 
|
124  | 
val _ = Context.reject_draft thy;  | 
|
125  | 
val cert = Thm.cterm_of thy;  | 
|
126  | 
val certT = Thm.ctyp_of thy;  | 
|
127  | 
||
| 
30473
 
e0b66c11e7e4
Assumption.all_prems_of, Assumption.all_assms_of;
 
wenzelm 
parents: 
29448 
diff
changeset
 | 
128  | 
val assms = Assumption.all_assms_of ctxt;  | 
| 28340 | 129  | 
val As = map Thm.term_of assms;  | 
130  | 
||
131  | 
val xs = map Free (fold Term.add_frees (prop :: As) []);  | 
|
132  | 
val fixes = map cert xs;  | 
|
133  | 
||
134  | 
val tfrees = fold Term.add_tfrees (prop :: As) [];  | 
|
135  | 
val instT = map (fn (a, S) => (certT (TVar ((a, 0), S)), certT (TFree (a, S)))) tfrees;  | 
|
136  | 
||
137  | 
val global_prop =  | 
|
| 
35845
 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 
wenzelm 
parents: 
35021 
diff
changeset
 | 
138  | 
cert (Term.map_types Logic.varifyT_global  | 
| 
 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 
wenzelm 
parents: 
35021 
diff
changeset
 | 
139  | 
(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
 | 
140  | 
|> Thm.weaken_sorts (Variable.sorts_of ctxt);  | 
| 
28973
 
c549650d1442
future proofs: pass actual futures to facilitate composite computations;
 
wenzelm 
parents: 
28619 
diff
changeset
 | 
141  | 
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
 | 
142  | 
(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
 | 
143  | 
Thm.adjust_maxidx_thm ~1 #>  | 
| 
28973
 
c549650d1442
future proofs: pass actual futures to facilitate composite computations;
 
wenzelm 
parents: 
28619 
diff
changeset
 | 
144  | 
Drule.implies_intr_list assms #>  | 
| 
 
c549650d1442
future proofs: pass actual futures to facilitate composite computations;
 
wenzelm 
parents: 
28619 
diff
changeset
 | 
145  | 
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
 | 
146  | 
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
 | 
147  | 
Thm.strip_shyps);  | 
| 28340 | 148  | 
val local_result =  | 
| 
32056
 
f4b74cbecdaf
future_result: explicitly impose Variable.sorts_of again;
 
wenzelm 
parents: 
30473 
diff
changeset
 | 
149  | 
Thm.future global_result global_prop  | 
| 28340 | 150  | 
|> Thm.instantiate (instT, [])  | 
151  | 
|> Drule.forall_elim_list fixes  | 
|
152  | 
|> fold (Thm.elim_implies o Thm.assume) assms;  | 
|
153  | 
in local_result end;  | 
|
154  | 
||
155  | 
||
| 
18027
 
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
 
wenzelm 
parents: 
17986 
diff
changeset
 | 
156  | 
|
| 
 
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
 
wenzelm 
parents: 
17986 
diff
changeset
 | 
157  | 
(** tactical theorem proving **)  | 
| 
 
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
 
wenzelm 
parents: 
17986 
diff
changeset
 | 
158  | 
|
| 23356 | 159  | 
(* prove_internal -- minimal checks, no normalization of result! *)  | 
| 
20250
 
c3f209752749
prove: proper assumption context, more tactic arguments;
 
wenzelm 
parents: 
20228 
diff
changeset
 | 
160  | 
|
| 23356 | 161  | 
fun prove_internal casms cprop tac =  | 
| 
20250
 
c3f209752749
prove: proper assumption context, more tactic arguments;
 
wenzelm 
parents: 
20228 
diff
changeset
 | 
162  | 
(case SINGLE (tac (map Assumption.assume casms)) (init cprop) of  | 
| 32197 | 163  | 
SOME th => Drule.implies_intr_list casms  | 
164  | 
(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
 | 
165  | 
| NONE => error "Tactic failed");  | 
| 
20250
 
c3f209752749
prove: proper assumption context, more tactic arguments;
 
wenzelm 
parents: 
20228 
diff
changeset
 | 
166  | 
|
| 
 
c3f209752749
prove: proper assumption context, more tactic arguments;
 
wenzelm 
parents: 
20228 
diff
changeset
 | 
167  | 
|
| 28340 | 168  | 
(* prove_common etc. *)  | 
| 17986 | 169  | 
|
| 28340 | 170  | 
fun prove_common immediate ctxt xs asms props tac =  | 
| 17980 | 171  | 
let  | 
| 21516 | 172  | 
val thy = ProofContext.theory_of ctxt;  | 
| 
26939
 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 
wenzelm 
parents: 
26713 
diff
changeset
 | 
173  | 
val string_of_term = Syntax.string_of_term ctxt;  | 
| 20056 | 174  | 
|
| 28355 | 175  | 
val pos = Position.thread_data ();  | 
| 
20250
 
c3f209752749
prove: proper assumption context, more tactic arguments;
 
wenzelm 
parents: 
20228 
diff
changeset
 | 
176  | 
fun err msg = cat_error msg  | 
| 
 
c3f209752749
prove: proper assumption context, more tactic arguments;
 
wenzelm 
parents: 
20228 
diff
changeset
 | 
177  | 
      ("The error(s) above occurred for the goal statement:\n" ^
 | 
| 28355 | 178  | 
string_of_term (Logic.list_implies (asms, Logic.mk_conjunction_list props)) ^  | 
179  | 
(case Position.str_of pos of "" => "" | s => "\n" ^ s));  | 
|
| 17980 | 180  | 
|
| 
20250
 
c3f209752749
prove: proper assumption context, more tactic arguments;
 
wenzelm 
parents: 
20228 
diff
changeset
 | 
181  | 
fun cert_safe t = Thm.cterm_of thy (Envir.beta_norm (Term.no_dummy_patterns t))  | 
| 17980 | 182  | 
handle TERM (msg, _) => err msg | TYPE (msg, _, _) => err msg;  | 
| 
20250
 
c3f209752749
prove: proper assumption context, more tactic arguments;
 
wenzelm 
parents: 
20228 
diff
changeset
 | 
183  | 
val casms = map cert_safe asms;  | 
| 
 
c3f209752749
prove: proper assumption context, more tactic arguments;
 
wenzelm 
parents: 
20228 
diff
changeset
 | 
184  | 
val cprops = map cert_safe props;  | 
| 17980 | 185  | 
|
| 
20250
 
c3f209752749
prove: proper assumption context, more tactic arguments;
 
wenzelm 
parents: 
20228 
diff
changeset
 | 
186  | 
val (prems, ctxt') = ctxt  | 
| 
 
c3f209752749
prove: proper assumption context, more tactic arguments;
 
wenzelm 
parents: 
20228 
diff
changeset
 | 
187  | 
|> Variable.add_fixes_direct xs  | 
| 
27218
 
4548c83cd508
prove: full Variable.declare_term, including constraints;
 
wenzelm 
parents: 
26939 
diff
changeset
 | 
188  | 
|> fold Variable.declare_term (asms @ props)  | 
| 
26713
 
1c6181def1d0
prove_global: Variable.set_body true, pass context;
 
wenzelm 
parents: 
26628 
diff
changeset
 | 
189  | 
|> Assumption.add_assumes casms  | 
| 
 
1c6181def1d0
prove_global: Variable.set_body true, pass context;
 
wenzelm 
parents: 
26628 
diff
changeset
 | 
190  | 
||> Variable.set_body true;  | 
| 
28619
 
89f9dd800a22
prove_common: include all sorts from context into statement, check shyps of result;
 
wenzelm 
parents: 
28446 
diff
changeset
 | 
191  | 
val sorts = Variable.sorts_of ctxt';  | 
| 17980 | 192  | 
|
| 
28619
 
89f9dd800a22
prove_common: include all sorts from context into statement, check shyps of result;
 
wenzelm 
parents: 
28446 
diff
changeset
 | 
193  | 
val stmt = Thm.weaken_sorts sorts (Conjunction.mk_conjunction_balanced cprops);  | 
| 28340 | 194  | 
|
195  | 
fun result () =  | 
|
196  | 
      (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
 | 
197  | 
NONE => err "Tactic failed"  | 
| 28340 | 198  | 
| SOME st =>  | 
| 32197 | 199  | 
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
 | 
200  | 
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
 | 
201  | 
then Thm.check_shyps sorts res  | 
| 28340 | 202  | 
            else err ("Proved a different theorem: " ^ string_of_term (Thm.prop_of res))
 | 
203  | 
end);  | 
|
204  | 
val res =  | 
|
| 
29448
 
34b9652b2f45
added Goal.future_enabled abstraction -- now also checks that this is already
 
wenzelm 
parents: 
29435 
diff
changeset
 | 
205  | 
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
 | 
206  | 
then result ()  | 
| 
37186
 
349e9223c685
explicit markup for forked goals, as indicated by Goal.fork;
 
wenzelm 
parents: 
36613 
diff
changeset
 | 
207  | 
else future_result ctxt' (fork result) (Thm.term_of stmt);  | 
| 17980 | 208  | 
in  | 
| 28340 | 209  | 
Conjunction.elim_balanced (length props) res  | 
| 20290 | 210  | 
|> map (Assumption.export false ctxt' ctxt)  | 
| 20056 | 211  | 
|> Variable.export ctxt' ctxt  | 
| 
20250
 
c3f209752749
prove: proper assumption context, more tactic arguments;
 
wenzelm 
parents: 
20228 
diff
changeset
 | 
212  | 
|> map Drule.zero_var_indexes  | 
| 17980 | 213  | 
end;  | 
214  | 
||
| 28341 | 215  | 
val prove_multi = prove_common true;  | 
| 17980 | 216  | 
|
| 
28446
 
a01de3b3fa2e
renamed promise to future, tuned related interfaces;
 
wenzelm 
parents: 
28363 
diff
changeset
 | 
217  | 
fun prove_future ctxt xs asms prop tac = hd (prove_common false ctxt xs asms [prop] tac);  | 
| 28340 | 218  | 
fun prove ctxt xs asms prop tac = hd (prove_common true ctxt xs asms [prop] tac);  | 
| 20056 | 219  | 
|
220  | 
fun prove_global thy xs asms prop tac =  | 
|
| 
36610
 
bafd82950e24
renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
 
wenzelm 
parents: 
35845 
diff
changeset
 | 
221  | 
Drule.export_without_context (prove (ProofContext.init_global thy) xs asms prop tac);  | 
| 
18027
 
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
 
wenzelm 
parents: 
17986 
diff
changeset
 | 
222  | 
|
| 
 
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
 
wenzelm 
parents: 
17986 
diff
changeset
 | 
223  | 
|
| 17980 | 224  | 
|
| 21687 | 225  | 
(** goal structure **)  | 
226  | 
||
227  | 
(* nested goals *)  | 
|
| 18207 | 228  | 
|
| 19184 | 229  | 
fun extract i n st =  | 
230  | 
(if i < 1 orelse n < 1 orelse i + n - 1 > Thm.nprems_of st then Seq.empty  | 
|
231  | 
else if n = 1 then Seq.single (Thm.cprem_of st i)  | 
|
| 23418 | 232  | 
else  | 
233  | 
Seq.single (Conjunction.mk_conjunction_balanced (map (Thm.cprem_of st) (i upto i + n - 1))))  | 
|
| 20260 | 234  | 
|> Seq.map (Thm.adjust_maxidx_cterm ~1 #> init);  | 
| 17980 | 235  | 
|
| 19221 | 236  | 
fun retrofit i n st' st =  | 
237  | 
(if n = 1 then st  | 
|
| 23418 | 238  | 
else st |> Drule.with_subgoal i (Conjunction.uncurry_balanced n))  | 
| 19221 | 239  | 
|> Thm.compose_no_flatten false (conclude st', Thm.nprems_of st') i;  | 
| 18207 | 240  | 
|
| 17980 | 241  | 
fun SELECT_GOAL tac i st =  | 
| 19191 | 242  | 
if Thm.nprems_of st = 1 andalso i = 1 then tac st  | 
| 19184 | 243  | 
else Seq.lifts (retrofit i 1) (Seq.maps tac (extract i 1 st)) st;  | 
| 17980 | 244  | 
|
| 21687 | 245  | 
|
246  | 
(* multiple goals *)  | 
|
247  | 
||
| 23418 | 248  | 
fun precise_conjunction_tac 0 i = eq_assume_tac i  | 
249  | 
| precise_conjunction_tac 1 i = SUBGOAL (K all_tac) i  | 
|
250  | 
| 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
 | 
251  | 
|
| 23418 | 252  | 
val adhoc_conjunction_tac = REPEAT_ALL_NEW  | 
253  | 
(SUBGOAL (fn (goal, i) =>  | 
|
254  | 
if can Logic.dest_conjunction goal then rtac Conjunction.conjunctionI i  | 
|
255  | 
else no_tac));  | 
|
| 21687 | 256  | 
|
| 23418 | 257  | 
val conjunction_tac = SUBGOAL (fn (goal, i) =>  | 
258  | 
precise_conjunction_tac (length (Logic.dest_conjunctions goal)) i ORELSE  | 
|
259  | 
TRY (adhoc_conjunction_tac i));  | 
|
| 21687 | 260  | 
|
| 23418 | 261  | 
val recover_conjunction_tac = PRIMITIVE (fn th =>  | 
262  | 
Conjunction.uncurry_balanced (Thm.nprems_of th) th);  | 
|
| 
23414
 
927203ad4b3a
tuned conjunction tactics: slightly smaller proof terms;
 
wenzelm 
parents: 
23356 
diff
changeset
 | 
263  | 
|
| 
 
927203ad4b3a
tuned conjunction tactics: slightly smaller proof terms;
 
wenzelm 
parents: 
23356 
diff
changeset
 | 
264  | 
fun PRECISE_CONJUNCTS n tac =  | 
| 
 
927203ad4b3a
tuned conjunction tactics: slightly smaller proof terms;
 
wenzelm 
parents: 
23356 
diff
changeset
 | 
265  | 
SELECT_GOAL (precise_conjunction_tac n 1  | 
| 
 
927203ad4b3a
tuned conjunction tactics: slightly smaller proof terms;
 
wenzelm 
parents: 
23356 
diff
changeset
 | 
266  | 
THEN tac  | 
| 23418 | 267  | 
THEN recover_conjunction_tac);  | 
| 
23414
 
927203ad4b3a
tuned conjunction tactics: slightly smaller proof terms;
 
wenzelm 
parents: 
23356 
diff
changeset
 | 
268  | 
|
| 21687 | 269  | 
fun CONJUNCTS tac =  | 
270  | 
SELECT_GOAL (conjunction_tac 1  | 
|
271  | 
THEN tac  | 
|
| 23418 | 272  | 
THEN recover_conjunction_tac);  | 
| 21687 | 273  | 
|
274  | 
||
275  | 
(* hhf normal form *)  | 
|
276  | 
||
277  | 
val norm_hhf_tac =  | 
|
278  | 
rtac Drule.asm_rl (*cheap approximation -- thanks to builtin Logic.flatten_params*)  | 
|
279  | 
THEN' SUBGOAL (fn (t, i) =>  | 
|
280  | 
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
 | 
281  | 
else rewrite_goal_tac Drule.norm_hhf_eqs i);  | 
| 21687 | 282  | 
|
| 25301 | 283  | 
fun compose_hhf_tac th i st =  | 
284  | 
PRIMSEQ (Thm.bicompose false (false, Drule.lift_all (Thm.cprem_of st i) th, 0) i) st;  | 
|
| 21687 | 285  | 
|
| 23237 | 286  | 
|
287  | 
(* non-atomic goal assumptions *)  | 
|
288  | 
||
| 23356 | 289  | 
fun non_atomic (Const ("==>", _) $ _ $ _) = true
 | 
290  | 
  | non_atomic (Const ("all", _) $ _) = true
 | 
|
291  | 
| non_atomic _ = false;  | 
|
292  | 
||
| 23237 | 293  | 
fun assume_rule_tac ctxt = norm_hhf_tac THEN' CSUBGOAL (fn (goal, i) =>  | 
294  | 
let  | 
|
295  | 
val ((_, goal'), ctxt') = Variable.focus goal ctxt;  | 
|
296  | 
val goal'' = Drule.cterm_rule (singleton (Variable.export ctxt' ctxt)) goal';  | 
|
| 23356 | 297  | 
val Rs = filter (non_atomic o Thm.term_of) (Drule.strip_imp_prems goal'');  | 
| 23237 | 298  | 
val tacs = Rs |> map (fn R =>  | 
| 
41228
 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 
wenzelm 
parents: 
39125 
diff
changeset
 | 
299  | 
Tactic.etac (Raw_Simplifier.norm_hhf (Thm.trivial R)) THEN_ALL_NEW assume_tac);  | 
| 23237 | 300  | 
in fold_rev (curry op APPEND') tacs (K no_tac) i end);  | 
301  | 
||
| 32365 | 302  | 
|
303  | 
(* parallel tacticals *)  | 
|
304  | 
||
305  | 
(*parallel choice of single results*)  | 
|
306  | 
fun PARALLEL_CHOICE tacs st =  | 
|
307  | 
(case Par_List.get_some (fn tac => SINGLE tac st) tacs of  | 
|
308  | 
NONE => Seq.empty  | 
|
309  | 
| SOME st' => Seq.single st');  | 
|
310  | 
||
311  | 
(*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
 | 
312  | 
exception FAILED of unit;  | 
| 32365 | 313  | 
fun PARALLEL_GOALS tac st =  | 
314  | 
let  | 
|
315  | 
val st0 = Thm.adjust_maxidx_thm ~1 st;  | 
|
316  | 
val _ = Thm.maxidx_of st0 >= 0 andalso  | 
|
317  | 
      raise THM ("PARALLEL_GOALS: schematic goal state", 0, [st]);
 | 
|
318  | 
||
319  | 
fun try_tac g =  | 
|
320  | 
(case SINGLE tac g of  | 
|
| 
32788
 
a65deb8f9434
PARALLEL_GOALS: proper scope for exception FAILED, with dummy argument to prevent its interpretation as variable;
 
wenzelm 
parents: 
32738 
diff
changeset
 | 
321  | 
NONE => raise FAILED ()  | 
| 32365 | 322  | 
| SOME g' => g');  | 
323  | 
||
324  | 
val goals = Drule.strip_imp_prems (Thm.cprop_of st0);  | 
|
325  | 
val results = Par_List.map (try_tac o init) goals;  | 
|
326  | 
in ALLGOALS (fn i => retrofit i 1 (nth results (i - 1))) st0 end  | 
|
| 
32788
 
a65deb8f9434
PARALLEL_GOALS: proper scope for exception FAILED, with dummy argument to prevent its interpretation as variable;
 
wenzelm 
parents: 
32738 
diff
changeset
 | 
327  | 
handle FAILED () => Seq.empty;  | 
| 32365 | 328  | 
|
| 18207 | 329  | 
end;  | 
330  | 
||
| 32365 | 331  | 
structure Basic_Goal: BASIC_GOAL = Goal;  | 
332  | 
open Basic_Goal;  |