| author | wenzelm | 
| Tue, 22 Oct 2024 12:45:38 +0200 | |
| changeset 81229 | e18600daa904 | 
| parent 79113 | 5109e4b2a292 | 
| child 82590 | d08f5b5ead0a | 
| permissions | -rw-r--r-- | 
| 17980 | 1 | (* Title: Pure/goal.ML | 
| 29345 | 2 | Author: Makarius | 
| 17980 | 3 | |
| 49061 
7449b804073b
central management of forked goals wrt. transaction id;
 wenzelm parents: 
49041diff
changeset | 4 | Goals in tactical theorem proving, with support for forked proofs. | 
| 17980 | 5 | *) | 
| 6 | ||
| 7 | signature BASIC_GOAL = | |
| 8 | sig | |
| 61527 | 9 | val quick_and_dirty: bool Config.T | 
| 17980 | 10 | val SELECT_GOAL: tactic -> int -> tactic | 
| 52458 
210bca64b894
less intrusive SELECT_GOAL: merely rearrange subgoals without detaching goal state, and thus preserve maxidx context;
 wenzelm parents: 
52456diff
changeset | 11 | val PREFER_GOAL: tactic -> int -> tactic | 
| 23418 | 12 | val CONJUNCTS: tactic -> int -> tactic | 
| 23414 
927203ad4b3a
tuned conjunction tactics: slightly smaller proof terms;
 wenzelm parents: 
23356diff
changeset | 13 | val PRECISE_CONJUNCTS: int -> tactic -> int -> tactic | 
| 17980 | 14 | end; | 
| 15 | ||
| 16 | signature GOAL = | |
| 17 | sig | |
| 18 | include BASIC_GOAL | |
| 19 | val init: cterm -> thm | |
| 52456 | 20 | val protect: int -> thm -> thm | 
| 17980 | 21 | val conclude: thm -> thm | 
| 49845 | 22 | val check_finished: Proof.context -> thm -> thm | 
| 32197 | 23 | val finish: Proof.context -> thm -> thm | 
| 54883 
dd04a8b654fc
proper context for norm_hhf and derived operations;
 wenzelm parents: 
54742diff
changeset | 24 | val norm_result: Proof.context -> thm -> thm | 
| 52499 | 25 | val skip_proofs_enabled: unit -> bool | 
| 28973 
c549650d1442
future proofs: pass actual futures to facilitate composite computations;
 wenzelm parents: 
28619diff
changeset | 26 | val future_result: Proof.context -> thm future -> term -> thm | 
| 74530 
823ccd84b879
revert bbfed17243af, breaks HOL-Proofs extraction;
 wenzelm parents: 
74526diff
changeset | 27 | val prove_internal: Proof.context -> cterm list -> cterm -> (thm list -> tactic) -> thm | 
| 59564 
fdc03c8daacc
Goal.prove_multi is superseded by the fully general Goal.prove_common;
 wenzelm parents: 
59498diff
changeset | 28 | val prove_common: Proof.context -> int option -> string list -> term list -> term list -> | 
| 74530 
823ccd84b879
revert bbfed17243af, breaks HOL-Proofs extraction;
 wenzelm parents: 
74526diff
changeset | 29 |     ({prems: thm list, context: Proof.context} -> tactic) -> thm list
 | 
| 28446 
a01de3b3fa2e
renamed promise to future, tuned related interfaces;
 wenzelm parents: 
28363diff
changeset | 30 | val prove_future: Proof.context -> string list -> term list -> term -> | 
| 74530 
823ccd84b879
revert bbfed17243af, breaks HOL-Proofs extraction;
 wenzelm parents: 
74526diff
changeset | 31 |     ({prems: thm list, context: Proof.context} -> tactic) -> thm
 | 
| 20290 | 32 | val prove: Proof.context -> string list -> term list -> term -> | 
| 74530 
823ccd84b879
revert bbfed17243af, breaks HOL-Proofs extraction;
 wenzelm parents: 
74526diff
changeset | 33 |     ({prems: thm list, context: Proof.context} -> tactic) -> thm
 | 
| 51118 
32a5994dd205
tuned signature -- legacy packages might want to fork proofs as well;
 wenzelm parents: 
51110diff
changeset | 34 | val prove_global_future: theory -> string list -> term list -> term -> | 
| 74530 
823ccd84b879
revert bbfed17243af, breaks HOL-Proofs extraction;
 wenzelm parents: 
74526diff
changeset | 35 |     ({prems: thm list, context: Proof.context} -> tactic) -> thm
 | 
| 26713 
1c6181def1d0
prove_global: Variable.set_body true, pass context;
 wenzelm parents: 
26628diff
changeset | 36 | val prove_global: theory -> string list -> term list -> term -> | 
| 74530 
823ccd84b879
revert bbfed17243af, breaks HOL-Proofs extraction;
 wenzelm parents: 
74526diff
changeset | 37 |     ({prems: thm list, context: Proof.context} -> tactic) -> thm
 | 
| 51551 | 38 | val prove_sorry: Proof.context -> string list -> term list -> term -> | 
| 74530 
823ccd84b879
revert bbfed17243af, breaks HOL-Proofs extraction;
 wenzelm parents: 
74526diff
changeset | 39 |     ({prems: thm list, context: Proof.context} -> tactic) -> thm
 | 
| 51551 | 40 | val prove_sorry_global: theory -> string list -> term list -> term -> | 
| 74530 
823ccd84b879
revert bbfed17243af, breaks HOL-Proofs extraction;
 wenzelm parents: 
74526diff
changeset | 41 |     ({prems: thm list, context: Proof.context} -> tactic) -> thm
 | 
| 52458 
210bca64b894
less intrusive SELECT_GOAL: merely rearrange subgoals without detaching goal state, and thus preserve maxidx context;
 wenzelm parents: 
52456diff
changeset | 42 | val restrict: int -> int -> thm -> thm | 
| 
210bca64b894
less intrusive SELECT_GOAL: merely rearrange subgoals without detaching goal state, and thus preserve maxidx context;
 wenzelm parents: 
52456diff
changeset | 43 | val unrestrict: int -> thm -> thm | 
| 23418 | 44 | val conjunction_tac: int -> tactic | 
| 23414 
927203ad4b3a
tuned conjunction tactics: slightly smaller proof terms;
 wenzelm parents: 
23356diff
changeset | 45 | val precise_conjunction_tac: int -> int -> tactic | 
| 23418 | 46 | val recover_conjunction_tac: tactic | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
54567diff
changeset | 47 | val norm_hhf_tac: Proof.context -> 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: 
17986diff
changeset | 54 | (** goals **) | 
| 
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
 wenzelm parents: 
17986diff
changeset | 55 | |
| 
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
 wenzelm parents: 
17986diff
changeset | 56 | (* | 
| 
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
 wenzelm parents: 
17986diff
changeset | 57 | -------- (init) | 
| 67721 | 58 | C \<Longrightarrow> #C | 
| 18027 
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
 wenzelm parents: 
17986diff
changeset | 59 | *) | 
| 77879 | 60 | fun init C = Thm.instantiate (TVars.empty, Vars.make1 ((("A", 0), propT), C)) Drule.protectI;
 | 
| 17980 | 61 | |
| 62 | (* | |
| 67721 | 63 | A1 \<Longrightarrow> ... \<Longrightarrow> An \<Longrightarrow> C | 
| 52456 | 64 | ------------------------ (protect n) | 
| 67721 | 65 | A1 \<Longrightarrow> ... \<Longrightarrow> An \<Longrightarrow> #C | 
| 17980 | 66 | *) | 
| 52456 | 67 | fun protect n th = Drule.comp_no_flatten (th, n) 1 Drule.protectI; | 
| 17980 | 68 | |
| 69 | (* | |
| 67721 | 70 | A \<Longrightarrow> ... \<Longrightarrow> #C | 
| 18027 
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
 wenzelm parents: 
17986diff
changeset | 71 | ---------------- (conclude) | 
| 67721 | 72 | A \<Longrightarrow> ... \<Longrightarrow> C | 
| 17980 | 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 = | 
| 51608 | 82 | if Thm.no_prems th then th | 
| 76051 | 83 |   else raise THM ("Proof failed.\n" ^ Goal_Display.string_of_goal ctxt th, 0, [th]);
 | 
| 32197 | 84 | |
| 49845 | 85 | fun finish ctxt = check_finished ctxt #> conclude; | 
| 17980 | 86 | |
| 87 | ||
| 18027 
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
 wenzelm parents: 
17986diff
changeset | 88 | |
| 
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
 wenzelm parents: 
17986diff
changeset | 89 | (** results **) | 
| 
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
 wenzelm parents: 
17986diff
changeset | 90 | |
| 28340 | 91 | (* normal form *) | 
| 92 | ||
| 54883 
dd04a8b654fc
proper context for norm_hhf and derived operations;
 wenzelm parents: 
54742diff
changeset | 93 | fun norm_result ctxt = | 
| 58950 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 wenzelm parents: 
58837diff
changeset | 94 | Drule.flexflex_unique (SOME ctxt) | 
| 54883 
dd04a8b654fc
proper context for norm_hhf and derived operations;
 wenzelm parents: 
54742diff
changeset | 95 | #> Raw_Simplifier.norm_hhf_protect ctxt | 
| 21604 
1af327306c8e
added norm/close_result (supercede local_standard etc.);
 wenzelm parents: 
21579diff
changeset | 96 | #> Thm.strip_shyps | 
| 
1af327306c8e
added norm/close_result (supercede local_standard etc.);
 wenzelm parents: 
21579diff
changeset | 97 | #> Drule.zero_var_indexes; | 
| 
1af327306c8e
added norm/close_result (supercede local_standard etc.);
 wenzelm parents: 
21579diff
changeset | 98 | |
| 
1af327306c8e
added norm/close_result (supercede local_standard etc.);
 wenzelm parents: 
21579diff
changeset | 99 | |
| 41703 
d27950860514
parallelization of nested Isar proofs is subject to Goal.parallel_proofs_threshold;
 wenzelm parents: 
41683diff
changeset | 100 | (* scheduling parameters *) | 
| 
d27950860514
parallelization of nested Isar proofs is subject to Goal.parallel_proofs_threshold;
 wenzelm parents: 
41683diff
changeset | 101 | |
| 52499 | 102 | fun skip_proofs_enabled () = | 
| 52710 
52790e3961fe
just one option "skip_proofs", without direct access within the editor (it makes document processing stateful without strong reasons -- monotonic updates already handle goal forks smoothly);
 wenzelm parents: 
52696diff
changeset | 103 | let val skip = Options.default_bool "skip_proofs" in | 
| 79113 
5109e4b2a292
pro-forma support for optional zproof: no proper content yet;
 wenzelm parents: 
78086diff
changeset | 104 | if Proofterm.any_proofs_enabled () andalso skip then | 
| 52499 | 105 | (warning "Proof terms enabled -- cannot skip proofs"; false) | 
| 106 | else skip | |
| 107 | end; | |
| 108 | ||
| 29448 
34b9652b2f45
added Goal.future_enabled abstraction -- now also checks that this is already
 wenzelm parents: 
29435diff
changeset | 109 | |
| 28446 
a01de3b3fa2e
renamed promise to future, tuned related interfaces;
 wenzelm parents: 
28363diff
changeset | 110 | (* future_result *) | 
| 28340 | 111 | |
| 28446 
a01de3b3fa2e
renamed promise to future, tuned related interfaces;
 wenzelm parents: 
28363diff
changeset | 112 | fun future_result ctxt result prop = | 
| 28340 | 113 | let | 
| 30473 
e0b66c11e7e4
Assumption.all_prems_of, Assumption.all_assms_of;
 wenzelm parents: 
29448diff
changeset | 114 | val assms = Assumption.all_assms_of ctxt; | 
| 28340 | 115 | val As = map Thm.term_of assms; | 
| 116 | ||
| 74266 | 117 | val frees = Frees.build (fold Frees.add_frees (prop :: As)); | 
| 118 | val xs = Frees.fold_rev (cons o Thm.cterm_of ctxt o Free o #1) frees []; | |
| 28340 | 119 | |
| 74266 | 120 | val tfrees = TFrees.build (fold TFrees.add_tfrees (prop :: As)); | 
| 121 | val Ts = Names.build (TFrees.fold (Names.add_set o #1 o #1) tfrees); | |
| 74248 | 122 | val instT = | 
| 74282 | 123 | TVars.build (tfrees |> TFrees.fold (fn ((a, S), _) => | 
| 124 | TVars.add (((a, 0), S), Thm.ctyp_of ctxt (TFree (a, S))))); | |
| 28340 | 125 | |
| 126 | val global_prop = | |
| 74248 | 127 | Logic.list_implies (As, prop) | 
| 74266 | 128 | |> Frees.fold_rev (Logic.all o Free o #1) frees | 
| 74248 | 129 | |> Logic.varify_types_global | 
| 59623 | 130 | |> Thm.cterm_of ctxt | 
| 61508 | 131 | |> Thm.weaken_sorts' ctxt; | 
| 28973 
c549650d1442
future proofs: pass actual futures to facilitate composite computations;
 wenzelm parents: 
28619diff
changeset | 132 | val global_result = result |> Future.map | 
| 58950 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 wenzelm parents: 
58837diff
changeset | 133 | (Drule.flexflex_unique (SOME ctxt) #> | 
| 28973 
c549650d1442
future proofs: pass actual futures to facilitate composite computations;
 wenzelm parents: 
28619diff
changeset | 134 | Drule.implies_intr_list assms #> | 
| 74248 | 135 | Drule.forall_intr_list xs #> | 
| 68693 
a9bef20b1e47
proper adjust_maxidx: assms could have maxidx >= 0, e.g. from command "subgoal premises";
 wenzelm parents: 
68130diff
changeset | 136 | Thm.adjust_maxidx_thm ~1 #> | 
| 74266 | 137 | Thm.generalize (Ts, Names.empty) 0 #> | 
| 70459 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 wenzelm parents: 
70398diff
changeset | 138 | Thm.strip_shyps #> | 
| 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 wenzelm parents: 
70398diff
changeset | 139 | Thm.solve_constraints); | 
| 28340 | 140 | val local_result = | 
| 32056 
f4b74cbecdaf
future_result: explicitly impose Variable.sorts_of again;
 wenzelm parents: 
30473diff
changeset | 141 | Thm.future global_result global_prop | 
| 70494 | 142 | |> Thm.close_derivation \<^here> | 
| 74282 | 143 | |> Thm.instantiate (instT, Vars.empty) | 
| 74248 | 144 | |> Drule.forall_elim_list xs | 
| 70459 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 wenzelm parents: 
70398diff
changeset | 145 | |> fold (Thm.elim_implies o Thm.assume) assms | 
| 
f0a445c5a82c
maintain sort constraints from type instantiations, with pro-forma derivation to collect oracles/thms;
 wenzelm parents: 
70398diff
changeset | 146 | |> Thm.solve_constraints; | 
| 28340 | 147 | in local_result end; | 
| 148 | ||
| 149 | ||
| 18027 
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
 wenzelm parents: 
17986diff
changeset | 150 | |
| 
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
 wenzelm parents: 
17986diff
changeset | 151 | (** tactical theorem proving **) | 
| 
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
 wenzelm parents: 
17986diff
changeset | 152 | |
| 23356 | 153 | (* prove_internal -- minimal checks, no normalization of result! *) | 
| 20250 
c3f209752749
prove: proper assumption context, more tactic arguments;
 wenzelm parents: 
20228diff
changeset | 154 | |
| 54883 
dd04a8b654fc
proper context for norm_hhf and derived operations;
 wenzelm parents: 
54742diff
changeset | 155 | fun prove_internal ctxt casms cprop tac = | 
| 74530 
823ccd84b879
revert bbfed17243af, breaks HOL-Proofs extraction;
 wenzelm parents: 
74526diff
changeset | 156 | (case SINGLE (tac (map (Assumption.assume ctxt) casms)) (init cprop) of | 
| 
823ccd84b879
revert bbfed17243af, breaks HOL-Proofs extraction;
 wenzelm parents: 
74526diff
changeset | 157 | SOME th => Drule.implies_intr_list casms (finish ctxt th) | 
| 
823ccd84b879
revert bbfed17243af, breaks HOL-Proofs extraction;
 wenzelm parents: 
74526diff
changeset | 158 | | NONE => error "Tactic failed"); | 
| 20250 
c3f209752749
prove: proper assumption context, more tactic arguments;
 wenzelm parents: 
20228diff
changeset | 159 | |
| 
c3f209752749
prove: proper assumption context, more tactic arguments;
 wenzelm parents: 
20228diff
changeset | 160 | |
| 51553 
63327f679cff
more ambitious Goal.skip_proofs: covers Goal.prove forms as well, and do not insist in quick_and_dirty (for the sake of Isabelle/jEdit);
 wenzelm parents: 
51552diff
changeset | 161 | (* prove variations *) | 
| 
63327f679cff
more ambitious Goal.skip_proofs: covers Goal.prove forms as well, and do not insist in quick_and_dirty (for the sake of Isabelle/jEdit);
 wenzelm parents: 
51552diff
changeset | 162 | |
| 59564 
fdc03c8daacc
Goal.prove_multi is superseded by the fully general Goal.prove_common;
 wenzelm parents: 
59498diff
changeset | 163 | fun prove_common ctxt fork_pri xs asms props tac = | 
| 17980 | 164 | let | 
| 42360 | 165 | val thy = Proof_Context.theory_of ctxt; | 
| 20056 | 166 | |
| 76047 | 167 | val schematic = exists Term.is_schematic props; | 
| 59564 
fdc03c8daacc
Goal.prove_multi is superseded by the fully general Goal.prove_common;
 wenzelm parents: 
59498diff
changeset | 168 | val immediate = is_none fork_pri; | 
| 79113 
5109e4b2a292
pro-forma support for optional zproof: no proper content yet;
 wenzelm parents: 
78086diff
changeset | 169 | val future = Future.proofs_enabled 1 andalso not (Proofterm.any_proofs_enabled ()); | 
| 52499 | 170 | val skip = not immediate andalso not schematic andalso future andalso skip_proofs_enabled (); | 
| 51553 
63327f679cff
more ambitious Goal.skip_proofs: covers Goal.prove forms as well, and do not insist in quick_and_dirty (for the sake of Isabelle/jEdit);
 wenzelm parents: 
51552diff
changeset | 171 | |
| 28355 | 172 | val pos = Position.thread_data (); | 
| 55559 | 173 | fun err msg = | 
| 174 | cat_error msg | |
| 175 |         ("The error(s) above occurred for the goal statement" ^ Position.here pos ^ ":\n" ^
 | |
| 176 | Syntax.string_of_term ctxt (Logic.list_implies (asms, Logic.mk_conjunction_list props))); | |
| 17980 | 177 | |
| 59623 | 178 | fun cert_safe t = Thm.cterm_of ctxt (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; | 
| 17980 | 188 | |
| 61508 | 189 | val stmt = Thm.weaken_sorts' ctxt' (Conjunction.mk_conjunction_balanced cprops); | 
| 28340 | 190 | |
| 51553 
63327f679cff
more ambitious Goal.skip_proofs: covers Goal.prove forms as well, and do not insist in quick_and_dirty (for the sake of Isabelle/jEdit);
 wenzelm parents: 
51552diff
changeset | 191 | fun tac' args st = | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
58963diff
changeset | 192 | if skip then ALLGOALS (Skip_Proof.cheat_tac ctxt) st before Skip_Proof.report ctxt | 
| 51553 
63327f679cff
more ambitious Goal.skip_proofs: covers Goal.prove forms as well, and do not insist in quick_and_dirty (for the sake of Isabelle/jEdit);
 wenzelm parents: 
51552diff
changeset | 193 | else tac args st; | 
| 28340 | 194 | fun result () = | 
| 51553 
63327f679cff
more ambitious Goal.skip_proofs: covers Goal.prove forms as well, and do not insist in quick_and_dirty (for the sake of Isabelle/jEdit);
 wenzelm parents: 
51552diff
changeset | 195 |       (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: 
38236diff
changeset | 196 | NONE => err "Tactic failed" | 
| 28340 | 197 | | SOME st => | 
| 54567 
cfe53047dc16
more uniform / rigid checking of Goal.prove_common vs. Proof.conclude_goal -- NB: Goal.prove_common cannot check hyps right now, e.g. due to undeclared Simplifier prems;
 wenzelm parents: 
53192diff
changeset | 198 | let | 
| 54981 | 199 | val _ = | 
| 65458 | 200 | Context.subthy_id (Thm.theory_id st, Context.theory_id thy) orelse | 
| 59582 | 201 | err "Bad background theory of goal state"; | 
| 54567 
cfe53047dc16
more uniform / rigid checking of Goal.prove_common vs. Proof.conclude_goal -- NB: Goal.prove_common cannot check hyps right now, e.g. due to undeclared Simplifier prems;
 wenzelm parents: 
53192diff
changeset | 202 | val res = | 
| 
cfe53047dc16
more uniform / rigid checking of Goal.prove_common vs. Proof.conclude_goal -- NB: Goal.prove_common cannot check hyps right now, e.g. due to undeclared Simplifier prems;
 wenzelm parents: 
53192diff
changeset | 203 | (finish ctxt' st | 
| 58950 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 wenzelm parents: 
58837diff
changeset | 204 | |> Drule.flexflex_unique (SOME ctxt') | 
| 61508 | 205 | |> Thm.check_shyps ctxt' | 
| 54994 
8e98d67325b1
reactivate Thm.check_hyps, after adapting AFP/Datatype_Order_Generator (see AFP/b7e389b5765c);
 wenzelm parents: 
54993diff
changeset | 206 | |> Thm.check_hyps (Context.Proof ctxt')) | 
| 54567 
cfe53047dc16
more uniform / rigid checking of Goal.prove_common vs. Proof.conclude_goal -- NB: Goal.prove_common cannot check hyps right now, e.g. due to undeclared Simplifier prems;
 wenzelm parents: 
53192diff
changeset | 207 | handle THM (msg, _, _) => err msg | ERROR msg => err msg; | 
| 
cfe53047dc16
more uniform / rigid checking of Goal.prove_common vs. Proof.conclude_goal -- NB: Goal.prove_common cannot check hyps right now, e.g. due to undeclared Simplifier prems;
 wenzelm parents: 
53192diff
changeset | 208 | in | 
| 76062 | 209 | if is_none (Unify.matcher (Context.Proof ctxt') [Thm.term_of stmt] [Thm.prop_of res]) | 
| 210 |             then err ("Proved a different theorem: " ^ Syntax.string_of_term ctxt' (Thm.prop_of res))
 | |
| 211 | else res | |
| 28340 | 212 | end); | 
| 213 | val res = | |
| 53192 | 214 | if immediate orelse schematic orelse not future orelse skip then result () | 
| 215 | else | |
| 216 | future_result ctxt' | |
| 59564 
fdc03c8daacc
Goal.prove_multi is superseded by the fully general Goal.prove_common;
 wenzelm parents: 
59498diff
changeset | 217 |           (Execution.fork {name = "Goal.prove", pos = Position.thread_data (), pri = the fork_pri}
 | 
| 
fdc03c8daacc
Goal.prove_multi is superseded by the fully general Goal.prove_common;
 wenzelm parents: 
59498diff
changeset | 218 | result) | 
| 53192 | 219 | (Thm.term_of stmt); | 
| 17980 | 220 | in | 
| 60722 
a8babbb6d5ea
normalize proof before splitting conjunctions, according to Proof.conclude_goal (see also 4dd0ba632e40) -- may reduce general resource usage;
 wenzelm parents: 
60695diff
changeset | 221 | res | 
| 70494 | 222 | |> Thm.close_derivation \<^here> | 
| 60722 
a8babbb6d5ea
normalize proof before splitting conjunctions, according to Proof.conclude_goal (see also 4dd0ba632e40) -- may reduce general resource usage;
 wenzelm parents: 
60695diff
changeset | 223 | |> Conjunction.elim_balanced (length props) | 
| 78086 | 224 | |> map (Assumption.export ctxt' ctxt) | 
| 20056 | 225 | |> Variable.export ctxt' ctxt | 
| 20250 
c3f209752749
prove: proper assumption context, more tactic arguments;
 wenzelm parents: 
20228diff
changeset | 226 | |> map Drule.zero_var_indexes | 
| 17980 | 227 | end; | 
| 228 | ||
| 59564 
fdc03c8daacc
Goal.prove_multi is superseded by the fully general Goal.prove_common;
 wenzelm parents: 
59498diff
changeset | 229 | fun prove_future_pri ctxt pri xs asms prop tac = | 
| 
fdc03c8daacc
Goal.prove_multi is superseded by the fully general Goal.prove_common;
 wenzelm parents: 
59498diff
changeset | 230 | hd (prove_common ctxt (SOME pri) xs asms [prop] tac); | 
| 17980 | 231 | |
| 59564 
fdc03c8daacc
Goal.prove_multi is superseded by the fully general Goal.prove_common;
 wenzelm parents: 
59498diff
changeset | 232 | fun prove_future ctxt = prove_future_pri ctxt ~1; | 
| 51118 
32a5994dd205
tuned signature -- legacy packages might want to fork proofs as well;
 wenzelm parents: 
51110diff
changeset | 233 | |
| 59564 
fdc03c8daacc
Goal.prove_multi is superseded by the fully general Goal.prove_common;
 wenzelm parents: 
59498diff
changeset | 234 | fun prove ctxt xs asms prop tac = hd (prove_common ctxt NONE xs asms [prop] tac); | 
| 20056 | 235 | |
| 51118 
32a5994dd205
tuned signature -- legacy packages might want to fork proofs as well;
 wenzelm parents: 
51110diff
changeset | 236 | fun prove_global_future thy xs asms prop tac = | 
| 
32a5994dd205
tuned signature -- legacy packages might want to fork proofs as well;
 wenzelm parents: 
51110diff
changeset | 237 | Drule.export_without_context (prove_future (Proof_Context.init_global thy) xs asms prop tac); | 
| 
32a5994dd205
tuned signature -- legacy packages might want to fork proofs as well;
 wenzelm parents: 
51110diff
changeset | 238 | |
| 20056 | 239 | fun prove_global thy xs asms prop tac = | 
| 42360 | 240 | 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: 
17986diff
changeset | 241 | |
| 61527 | 242 | |
| 243 | (* skip proofs *) | |
| 244 | ||
| 69575 | 245 | val quick_and_dirty = Config.declare_option_bool ("quick_and_dirty", \<^here>);
 | 
| 61527 | 246 | |
| 51551 | 247 | fun prove_sorry ctxt xs asms prop tac = | 
| 52059 | 248 | if Config.get ctxt quick_and_dirty then | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
58963diff
changeset | 249 | prove ctxt xs asms prop (fn _ => ALLGOALS (Skip_Proof.cheat_tac ctxt)) | 
| 68130 
6fb85346cb79
clarified future scheduling parameters, with support for parallel_limit;
 wenzelm parents: 
68025diff
changeset | 250 | else (if Future.proofs_enabled 1 then prove_future_pri ctxt ~2 else prove ctxt) xs asms prop tac; | 
| 51551 | 251 | |
| 252 | fun prove_sorry_global thy xs asms prop tac = | |
| 253 | Drule.export_without_context | |
| 254 | (prove_sorry (Proof_Context.init_global thy) xs asms prop tac); | |
| 255 | ||
| 18027 
09ab79d4e8e1
renamed Goal constant to prop, more general protect/unprotect interfaces;
 wenzelm parents: 
17986diff
changeset | 256 | |
| 17980 | 257 | |
| 21687 | 258 | (** goal structure **) | 
| 259 | ||
| 52458 
210bca64b894
less intrusive SELECT_GOAL: merely rearrange subgoals without detaching goal state, and thus preserve maxidx context;
 wenzelm parents: 
52456diff
changeset | 260 | (* rearrange subgoals *) | 
| 
210bca64b894
less intrusive SELECT_GOAL: merely rearrange subgoals without detaching goal state, and thus preserve maxidx context;
 wenzelm parents: 
52456diff
changeset | 261 | |
| 
210bca64b894
less intrusive SELECT_GOAL: merely rearrange subgoals without detaching goal state, and thus preserve maxidx context;
 wenzelm parents: 
52456diff
changeset | 262 | fun restrict i n st = | 
| 
210bca64b894
less intrusive SELECT_GOAL: merely rearrange subgoals without detaching goal state, and thus preserve maxidx context;
 wenzelm parents: 
52456diff
changeset | 263 | if i < 1 orelse n < 1 orelse i + n - 1 > Thm.nprems_of st | 
| 
210bca64b894
less intrusive SELECT_GOAL: merely rearrange subgoals without detaching goal state, and thus preserve maxidx context;
 wenzelm parents: 
52456diff
changeset | 264 |   then raise THM ("Goal.restrict", i, [st])
 | 
| 
210bca64b894
less intrusive SELECT_GOAL: merely rearrange subgoals without detaching goal state, and thus preserve maxidx context;
 wenzelm parents: 
52456diff
changeset | 265 | else rotate_prems (i - 1) st |> protect n; | 
| 
210bca64b894
less intrusive SELECT_GOAL: merely rearrange subgoals without detaching goal state, and thus preserve maxidx context;
 wenzelm parents: 
52456diff
changeset | 266 | |
| 
210bca64b894
less intrusive SELECT_GOAL: merely rearrange subgoals without detaching goal state, and thus preserve maxidx context;
 wenzelm parents: 
52456diff
changeset | 267 | fun unrestrict i = conclude #> rotate_prems (1 - i); | 
| 
210bca64b894
less intrusive SELECT_GOAL: merely rearrange subgoals without detaching goal state, and thus preserve maxidx context;
 wenzelm parents: 
52456diff
changeset | 268 | |
| 
210bca64b894
less intrusive SELECT_GOAL: merely rearrange subgoals without detaching goal state, and thus preserve maxidx context;
 wenzelm parents: 
52456diff
changeset | 269 | (*with structural marker*) | 
| 17980 | 270 | fun SELECT_GOAL tac i st = | 
| 19191 | 271 | if Thm.nprems_of st = 1 andalso i = 1 then tac st | 
| 52458 
210bca64b894
less intrusive SELECT_GOAL: merely rearrange subgoals without detaching goal state, and thus preserve maxidx context;
 wenzelm parents: 
52456diff
changeset | 272 | else (PRIMITIVE (restrict i 1) THEN tac THEN PRIMITIVE (unrestrict i)) st; | 
| 
210bca64b894
less intrusive SELECT_GOAL: merely rearrange subgoals without detaching goal state, and thus preserve maxidx context;
 wenzelm parents: 
52456diff
changeset | 273 | |
| 
210bca64b894
less intrusive SELECT_GOAL: merely rearrange subgoals without detaching goal state, and thus preserve maxidx context;
 wenzelm parents: 
52456diff
changeset | 274 | (*without structural marker*) | 
| 
210bca64b894
less intrusive SELECT_GOAL: merely rearrange subgoals without detaching goal state, and thus preserve maxidx context;
 wenzelm parents: 
52456diff
changeset | 275 | fun PREFER_GOAL tac i st = | 
| 
210bca64b894
less intrusive SELECT_GOAL: merely rearrange subgoals without detaching goal state, and thus preserve maxidx context;
 wenzelm parents: 
52456diff
changeset | 276 | if i < 1 orelse i > Thm.nprems_of st then Seq.empty | 
| 
210bca64b894
less intrusive SELECT_GOAL: merely rearrange subgoals without detaching goal state, and thus preserve maxidx context;
 wenzelm parents: 
52456diff
changeset | 277 | else (PRIMITIVE (rotate_prems (i - 1)) THEN tac THEN PRIMITIVE (rotate_prems (1 - i))) st; | 
| 17980 | 278 | |
| 21687 | 279 | |
| 280 | (* multiple goals *) | |
| 281 | ||
| 23418 | 282 | fun precise_conjunction_tac 0 i = eq_assume_tac i | 
| 283 | | precise_conjunction_tac 1 i = SUBGOAL (K all_tac) i | |
| 284 | | 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 | 285 | |
| 23418 | 286 | val adhoc_conjunction_tac = REPEAT_ALL_NEW | 
| 287 | (SUBGOAL (fn (goal, i) => | |
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
58963diff
changeset | 288 | if can Logic.dest_conjunction goal then resolve0_tac [Conjunction.conjunctionI] i | 
| 23418 | 289 | else no_tac)); | 
| 21687 | 290 | |
| 23418 | 291 | val conjunction_tac = SUBGOAL (fn (goal, i) => | 
| 292 | precise_conjunction_tac (length (Logic.dest_conjunctions goal)) i ORELSE | |
| 293 | TRY (adhoc_conjunction_tac i)); | |
| 21687 | 294 | |
| 23418 | 295 | val recover_conjunction_tac = PRIMITIVE (fn th => | 
| 296 | Conjunction.uncurry_balanced (Thm.nprems_of th) th); | |
| 23414 
927203ad4b3a
tuned conjunction tactics: slightly smaller proof terms;
 wenzelm parents: 
23356diff
changeset | 297 | |
| 
927203ad4b3a
tuned conjunction tactics: slightly smaller proof terms;
 wenzelm parents: 
23356diff
changeset | 298 | fun PRECISE_CONJUNCTS n tac = | 
| 
927203ad4b3a
tuned conjunction tactics: slightly smaller proof terms;
 wenzelm parents: 
23356diff
changeset | 299 | SELECT_GOAL (precise_conjunction_tac n 1 | 
| 
927203ad4b3a
tuned conjunction tactics: slightly smaller proof terms;
 wenzelm parents: 
23356diff
changeset | 300 | THEN tac | 
| 23418 | 301 | THEN recover_conjunction_tac); | 
| 23414 
927203ad4b3a
tuned conjunction tactics: slightly smaller proof terms;
 wenzelm parents: 
23356diff
changeset | 302 | |
| 21687 | 303 | fun CONJUNCTS tac = | 
| 304 | SELECT_GOAL (conjunction_tac 1 | |
| 305 | THEN tac | |
| 23418 | 306 | THEN recover_conjunction_tac); | 
| 21687 | 307 | |
| 308 | ||
| 309 | (* hhf normal form *) | |
| 310 | ||
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
54567diff
changeset | 311 | fun norm_hhf_tac ctxt = | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
58963diff
changeset | 312 | resolve_tac ctxt [Drule.asm_rl] (*cheap approximation -- thanks to builtin Logic.flatten_params*) | 
| 21687 | 313 | THEN' SUBGOAL (fn (t, i) => | 
| 78046 
78deba4fdf27
tuned: more accurate check (is_norm_hhf protect);
 wenzelm parents: 
77879diff
changeset | 314 |     if Drule.is_norm_hhf {protect = false} t then all_tac
 | 
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
54567diff
changeset | 315 | else rewrite_goal_tac ctxt Drule.norm_hhf_eqs i); | 
| 21687 | 316 | |
| 23237 | 317 | |
| 318 | (* non-atomic goal assumptions *) | |
| 319 | ||
| 56245 | 320 | fun non_atomic (Const ("Pure.imp", _) $ _ $ _) = true
 | 
| 321 |   | non_atomic (Const ("Pure.all", _) $ _) = true
 | |
| 23356 | 322 | | non_atomic _ = false; | 
| 323 | ||
| 54742 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 wenzelm parents: 
54567diff
changeset | 324 | fun assume_rule_tac ctxt = norm_hhf_tac ctxt THEN' CSUBGOAL (fn (goal, i) => | 
| 23237 | 325 | let | 
| 60695 
757549b4bbe6
Variable.focus etc.: optional bindings provided by user;
 wenzelm parents: 
60642diff
changeset | 326 | val ((_, goal'), ctxt') = Variable.focus_cterm NONE goal ctxt; | 
| 23237 | 327 | val goal'' = Drule.cterm_rule (singleton (Variable.export ctxt' ctxt)) goal'; | 
| 23356 | 328 | val Rs = filter (non_atomic o Thm.term_of) (Drule.strip_imp_prems goal''); | 
| 23237 | 329 | val tacs = Rs |> map (fn R => | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
58963diff
changeset | 330 | eresolve_tac ctxt [Raw_Simplifier.norm_hhf ctxt (Thm.trivial R)] THEN_ALL_NEW assume_tac ctxt); | 
| 23237 | 331 | in fold_rev (curry op APPEND') tacs (K no_tac) i end); | 
| 332 | ||
| 52461 
ef4c16887f83
more scalable PARALLEL_GOALS, using more elementary retrofit;
 wenzelm parents: 
52458diff
changeset | 333 | end; | 
| 
ef4c16887f83
more scalable PARALLEL_GOALS, using more elementary retrofit;
 wenzelm parents: 
52458diff
changeset | 334 | |
| 32365 | 335 | structure Basic_Goal: BASIC_GOAL = Goal; | 
| 336 | open Basic_Goal; |