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