src/Pure/goal.ML
author wenzelm
Sat, 29 May 2010 19:46:29 +0200
changeset 37186 349e9223c685
parent 36613 f3157c288aca
child 37690 b16231572c61
permissions -rw-r--r--
explicit markup for forked goals, as indicated by Goal.fork; accumulate pending forks within command state and hilight accordingly; Isabelle_Process: enforce future_terminal_proof, which gives some impression of non-linear/parallel checking;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/goal.ML
29345
5904873d8f11 tuned protect, conclude: Drule.comp_no_flatten;
wenzelm
parents: 29343
diff changeset
     2
    Author:     Makarius
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
     3
18139
wenzelm
parents: 18119
diff changeset
     4
Goals in tactical theorem proving.
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
     5
*)
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
     6
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
     7
signature BASIC_GOAL =
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
     8
sig
32738
15bb09ca0378 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32365
diff changeset
     9
  val parallel_proofs: int Unsynchronized.ref
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    10
  val SELECT_GOAL: tactic -> int -> tactic
23418
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
    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
9b74d0339c44 added PARALLEL_CHOICE, PARALLEL_GOALS;
wenzelm
parents: 32255
diff changeset
    13
  val PARALLEL_CHOICE: tactic list -> tactic
9b74d0339c44 added PARALLEL_CHOICE, PARALLEL_GOALS;
wenzelm
parents: 32255
diff changeset
    14
  val PARALLEL_GOALS: tactic -> tactic
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    15
end;
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    16
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    17
signature GOAL =
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    18
sig
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    19
  include BASIC_GOAL
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    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
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    22
  val conclude: thm -> thm
32197
bc341bbe4417 Goal.finish: explicit context for printing;
wenzelm
parents: 32187
diff changeset
    23
  val check_finished: Proof.context -> thm -> unit
bc341bbe4417 Goal.finish: explicit context for printing;
wenzelm
parents: 32187
diff changeset
    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
37186
349e9223c685 explicit markup for forked goals, as indicated by Goal.fork;
wenzelm
parents: 36613
diff changeset
    29
  val local_future_enforced: unit -> bool
28973
c549650d1442 future proofs: pass actual futures to facilitate composite computations;
wenzelm
parents: 28619
diff changeset
    30
  val future_result: Proof.context -> thm future -> term -> thm
23356
dbe3731241c3 renamed prove_raw to prove_internal;
wenzelm
parents: 23237
diff changeset
    31
  val prove_internal: cterm list -> cterm -> (thm list -> tactic) -> thm
20290
3f886c176c9b normalized Proof.context/method type aliases;
wenzelm
parents: 20260
diff changeset
    32
  val prove_multi: Proof.context -> string list -> term list -> term list ->
3f886c176c9b normalized Proof.context/method type aliases;
wenzelm
parents: 20260
diff changeset
    33
    ({prems: thm list, context: Proof.context} -> tactic) -> thm list
28446
a01de3b3fa2e renamed promise to future, tuned related interfaces;
wenzelm
parents: 28363
diff changeset
    34
  val prove_future: Proof.context -> string list -> term list -> term ->
28340
e8597242f649 added promise_result, prove_promise;
wenzelm
parents: 27218
diff changeset
    35
    ({prems: thm list, context: Proof.context} -> tactic) -> thm
20290
3f886c176c9b normalized Proof.context/method type aliases;
wenzelm
parents: 20260
diff changeset
    36
  val prove: Proof.context -> string list -> term list -> term ->
3f886c176c9b normalized Proof.context/method type aliases;
wenzelm
parents: 20260
diff changeset
    37
    ({prems: thm list, context: Proof.context} -> tactic) -> thm
26713
1c6181def1d0 prove_global: Variable.set_body true, pass context;
wenzelm
parents: 26628
diff changeset
    38
  val prove_global: theory -> string list -> term list -> term ->
1c6181def1d0 prove_global: Variable.set_body true, pass context;
wenzelm
parents: 26628
diff changeset
    39
    ({prems: thm list, context: Proof.context} -> tactic) -> thm
19184
3e30297e1300 added extract, retrofit;
wenzelm
parents: 18678
diff changeset
    40
  val extract: int -> int -> thm -> thm Seq.seq
3e30297e1300 added extract, retrofit;
wenzelm
parents: 18678
diff changeset
    41
  val retrofit: int -> int -> thm -> thm -> thm Seq.seq
23418
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
    42
  val conjunction_tac: int -> tactic
23414
927203ad4b3a tuned conjunction tactics: slightly smaller proof terms;
wenzelm
parents: 23356
diff changeset
    43
  val precise_conjunction_tac: int -> int -> tactic
23418
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
    44
  val recover_conjunction_tac: tactic
21687
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
    45
  val norm_hhf_tac: int -> tactic
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
    46
  val compose_hhf_tac: thm -> int -> tactic
23237
ac9d126456e1 added assume_rule_tac;
wenzelm
parents: 22902
diff changeset
    47
  val assume_rule_tac: Proof.context -> int -> tactic
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    48
end;
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    49
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    50
structure Goal: GOAL =
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    51
struct
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    52
18027
09ab79d4e8e1 renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17986
diff changeset
    53
(** goals **)
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
(*
09ab79d4e8e1 renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17986
diff changeset
    56
  -------- (init)
09ab79d4e8e1 renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17986
diff changeset
    57
  C ==> #C
09ab79d4e8e1 renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17986
diff changeset
    58
*)
20290
3f886c176c9b normalized Proof.context/method type aliases;
wenzelm
parents: 20260
diff changeset
    59
val init =
22902
ac833b4bb7ee moved some Drule operations to Thm (see more_thm.ML);
wenzelm
parents: 21687
diff changeset
    60
  let val A = #1 (Thm.dest_implies (Thm.cprop_of Drule.protectI))
20290
3f886c176c9b normalized Proof.context/method type aliases;
wenzelm
parents: 20260
diff changeset
    61
  in fn C => Thm.instantiate ([], [(A, C)]) Drule.protectI end;
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    62
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    63
(*
18119
63901a9b99dc export compose_hhf;
wenzelm
parents: 18027
diff changeset
    64
   C
63901a9b99dc export compose_hhf;
wenzelm
parents: 18027
diff changeset
    65
  --- (protect)
63901a9b99dc export compose_hhf;
wenzelm
parents: 18027
diff changeset
    66
  #C
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    67
*)
29345
5904873d8f11 tuned protect, conclude: Drule.comp_no_flatten;
wenzelm
parents: 29343
diff changeset
    68
fun protect th = Drule.comp_no_flatten (th, 0) 1 Drule.protectI;
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    69
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    70
(*
18027
09ab79d4e8e1 renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17986
diff changeset
    71
  A ==> ... ==> #C
09ab79d4e8e1 renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17986
diff changeset
    72
  ---------------- (conclude)
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    73
  A ==> ... ==> C
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    74
*)
29345
5904873d8f11 tuned protect, conclude: Drule.comp_no_flatten;
wenzelm
parents: 29343
diff changeset
    75
fun conclude th = Drule.comp_no_flatten (th, Thm.nprems_of th) 1 Drule.protectD;
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    76
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    77
(*
18027
09ab79d4e8e1 renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17986
diff changeset
    78
  #C
09ab79d4e8e1 renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17986
diff changeset
    79
  --- (finish)
09ab79d4e8e1 renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17986
diff changeset
    80
   C
17983
wenzelm
parents: 17980
diff changeset
    81
*)
32197
bc341bbe4417 Goal.finish: explicit context for printing;
wenzelm
parents: 32187
diff changeset
    82
fun check_finished ctxt th =
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    83
  (case Thm.nprems_of th of
32197
bc341bbe4417 Goal.finish: explicit context for printing;
wenzelm
parents: 32187
diff changeset
    84
    0 => ()
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    85
  | n => raise THM ("Proof failed.\n" ^
32197
bc341bbe4417 Goal.finish: explicit context for printing;
wenzelm
parents: 32187
diff changeset
    86
      Pretty.string_of (Pretty.chunks
bc341bbe4417 Goal.finish: explicit context for printing;
wenzelm
parents: 32187
diff changeset
    87
        (Goal_Display.pretty_goals ctxt {total = true, main = true, maxgoals = n} th)) ^
bc341bbe4417 Goal.finish: explicit context for printing;
wenzelm
parents: 32187
diff changeset
    88
      "\n" ^ string_of_int n ^ " unsolved goal(s)!", 0, [th]));
bc341bbe4417 Goal.finish: explicit context for printing;
wenzelm
parents: 32187
diff changeset
    89
bc341bbe4417 Goal.finish: explicit context for printing;
wenzelm
parents: 32187
diff changeset
    90
fun finish ctxt th = (check_finished ctxt th; conclude th);
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    91
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    92
18027
09ab79d4e8e1 renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17986
diff changeset
    93
09ab79d4e8e1 renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17986
diff changeset
    94
(** results **)
09ab79d4e8e1 renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17986
diff changeset
    95
28340
e8597242f649 added promise_result, prove_promise;
wenzelm
parents: 27218
diff changeset
    96
(* normal form *)
e8597242f649 added promise_result, prove_promise;
wenzelm
parents: 27218
diff changeset
    97
21604
1af327306c8e added norm/close_result (supercede local_standard etc.);
wenzelm
parents: 21579
diff changeset
    98
val norm_result =
1af327306c8e added norm/close_result (supercede local_standard etc.);
wenzelm
parents: 21579
diff changeset
    99
  Drule.flexflex_unique
1af327306c8e added norm/close_result (supercede local_standard etc.);
wenzelm
parents: 21579
diff changeset
   100
  #> MetaSimplifier.norm_hhf_protect
1af327306c8e added norm/close_result (supercede local_standard etc.);
wenzelm
parents: 21579
diff changeset
   101
  #> Thm.strip_shyps
1af327306c8e added norm/close_result (supercede local_standard etc.);
wenzelm
parents: 21579
diff changeset
   102
  #> Drule.zero_var_indexes;
1af327306c8e added norm/close_result (supercede local_standard etc.);
wenzelm
parents: 21579
diff changeset
   103
1af327306c8e added norm/close_result (supercede local_standard etc.);
wenzelm
parents: 21579
diff changeset
   104
37186
349e9223c685 explicit markup for forked goals, as indicated by Goal.fork;
wenzelm
parents: 36613
diff changeset
   105
(* parallel proofs *)
349e9223c685 explicit markup for forked goals, as indicated by Goal.fork;
wenzelm
parents: 36613
diff changeset
   106
349e9223c685 explicit markup for forked goals, as indicated by Goal.fork;
wenzelm
parents: 36613
diff changeset
   107
fun fork e = Future.fork_pri ~1 (fn () =>
349e9223c685 explicit markup for forked goals, as indicated by Goal.fork;
wenzelm
parents: 36613
diff changeset
   108
  let
349e9223c685 explicit markup for forked goals, as indicated by Goal.fork;
wenzelm
parents: 36613
diff changeset
   109
    val _ = Output.status (Markup.markup Markup.forked "");
349e9223c685 explicit markup for forked goals, as indicated by Goal.fork;
wenzelm
parents: 36613
diff changeset
   110
    val x = e ();  (*sic*)
349e9223c685 explicit markup for forked goals, as indicated by Goal.fork;
wenzelm
parents: 36613
diff changeset
   111
    val _ = Output.status (Markup.markup Markup.joined "");
349e9223c685 explicit markup for forked goals, as indicated by Goal.fork;
wenzelm
parents: 36613
diff changeset
   112
  in x end);
29448
34b9652b2f45 added Goal.future_enabled abstraction -- now also checks that this is already
wenzelm
parents: 29435
diff changeset
   113
32738
15bb09ca0378 explicit indication of Unsynchronized.ref;
wenzelm
parents: 32365
diff changeset
   114
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
   115
34b9652b2f45 added Goal.future_enabled abstraction -- now also checks that this is already
wenzelm
parents: 29435
diff changeset
   116
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
   117
  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
   118
11f8ee55662d parallel_proofs: more fine-grained control with optional parallel checking of nested Isar proofs;
wenzelm
parents: 32058
diff changeset
   119
fun local_future_enabled () = future_enabled () andalso ! parallel_proofs >= 2;
37186
349e9223c685 explicit markup for forked goals, as indicated by Goal.fork;
wenzelm
parents: 36613
diff changeset
   120
fun local_future_enforced () = future_enabled () andalso ! parallel_proofs >= 3;
29448
34b9652b2f45 added Goal.future_enabled abstraction -- now also checks that this is already
wenzelm
parents: 29435
diff changeset
   121
34b9652b2f45 added Goal.future_enabled abstraction -- now also checks that this is already
wenzelm
parents: 29435
diff changeset
   122
28446
a01de3b3fa2e renamed promise to future, tuned related interfaces;
wenzelm
parents: 28363
diff changeset
   123
(* future_result *)
28340
e8597242f649 added promise_result, prove_promise;
wenzelm
parents: 27218
diff changeset
   124
28446
a01de3b3fa2e renamed promise to future, tuned related interfaces;
wenzelm
parents: 28363
diff changeset
   125
fun future_result ctxt result prop =
28340
e8597242f649 added promise_result, prove_promise;
wenzelm
parents: 27218
diff changeset
   126
  let
e8597242f649 added promise_result, prove_promise;
wenzelm
parents: 27218
diff changeset
   127
    val thy = ProofContext.theory_of ctxt;
e8597242f649 added promise_result, prove_promise;
wenzelm
parents: 27218
diff changeset
   128
    val _ = Context.reject_draft thy;
e8597242f649 added promise_result, prove_promise;
wenzelm
parents: 27218
diff changeset
   129
    val cert = Thm.cterm_of thy;
e8597242f649 added promise_result, prove_promise;
wenzelm
parents: 27218
diff changeset
   130
    val certT = Thm.ctyp_of thy;
e8597242f649 added promise_result, prove_promise;
wenzelm
parents: 27218
diff changeset
   131
30473
e0b66c11e7e4 Assumption.all_prems_of, Assumption.all_assms_of;
wenzelm
parents: 29448
diff changeset
   132
    val assms = Assumption.all_assms_of ctxt;
28340
e8597242f649 added promise_result, prove_promise;
wenzelm
parents: 27218
diff changeset
   133
    val As = map Thm.term_of assms;
e8597242f649 added promise_result, prove_promise;
wenzelm
parents: 27218
diff changeset
   134
e8597242f649 added promise_result, prove_promise;
wenzelm
parents: 27218
diff changeset
   135
    val xs = map Free (fold Term.add_frees (prop :: As) []);
e8597242f649 added promise_result, prove_promise;
wenzelm
parents: 27218
diff changeset
   136
    val fixes = map cert xs;
e8597242f649 added promise_result, prove_promise;
wenzelm
parents: 27218
diff changeset
   137
e8597242f649 added promise_result, prove_promise;
wenzelm
parents: 27218
diff changeset
   138
    val tfrees = fold Term.add_tfrees (prop :: As) [];
e8597242f649 added promise_result, prove_promise;
wenzelm
parents: 27218
diff changeset
   139
    val instT = map (fn (a, S) => (certT (TVar ((a, 0), S)), certT (TFree (a, S)))) tfrees;
e8597242f649 added promise_result, prove_promise;
wenzelm
parents: 27218
diff changeset
   140
e8597242f649 added promise_result, prove_promise;
wenzelm
parents: 27218
diff changeset
   141
    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
   142
      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
   143
        (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
   144
      |> Thm.weaken_sorts (Variable.sorts_of ctxt);
28973
c549650d1442 future proofs: pass actual futures to facilitate composite computations;
wenzelm
parents: 28619
diff changeset
   145
    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
   146
      (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
   147
        Thm.adjust_maxidx_thm ~1 #>
28973
c549650d1442 future proofs: pass actual futures to facilitate composite computations;
wenzelm
parents: 28619
diff changeset
   148
        Drule.implies_intr_list assms #>
c549650d1442 future proofs: pass actual futures to facilitate composite computations;
wenzelm
parents: 28619
diff changeset
   149
        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
   150
        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
   151
        Thm.strip_shyps);
28340
e8597242f649 added promise_result, prove_promise;
wenzelm
parents: 27218
diff changeset
   152
    val local_result =
32056
f4b74cbecdaf future_result: explicitly impose Variable.sorts_of again;
wenzelm
parents: 30473
diff changeset
   153
      Thm.future global_result global_prop
28340
e8597242f649 added promise_result, prove_promise;
wenzelm
parents: 27218
diff changeset
   154
      |> Thm.instantiate (instT, [])
e8597242f649 added promise_result, prove_promise;
wenzelm
parents: 27218
diff changeset
   155
      |> Drule.forall_elim_list fixes
e8597242f649 added promise_result, prove_promise;
wenzelm
parents: 27218
diff changeset
   156
      |> fold (Thm.elim_implies o Thm.assume) assms;
e8597242f649 added promise_result, prove_promise;
wenzelm
parents: 27218
diff changeset
   157
  in local_result end;
e8597242f649 added promise_result, prove_promise;
wenzelm
parents: 27218
diff changeset
   158
e8597242f649 added promise_result, prove_promise;
wenzelm
parents: 27218
diff changeset
   159
18027
09ab79d4e8e1 renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17986
diff changeset
   160
09ab79d4e8e1 renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17986
diff changeset
   161
(** tactical theorem proving **)
09ab79d4e8e1 renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17986
diff changeset
   162
23356
dbe3731241c3 renamed prove_raw to prove_internal;
wenzelm
parents: 23237
diff changeset
   163
(* prove_internal -- minimal checks, no normalization of result! *)
20250
c3f209752749 prove: proper assumption context, more tactic arguments;
wenzelm
parents: 20228
diff changeset
   164
23356
dbe3731241c3 renamed prove_raw to prove_internal;
wenzelm
parents: 23237
diff changeset
   165
fun prove_internal casms cprop tac =
20250
c3f209752749 prove: proper assumption context, more tactic arguments;
wenzelm
parents: 20228
diff changeset
   166
  (case SINGLE (tac (map Assumption.assume casms)) (init cprop) of
32197
bc341bbe4417 Goal.finish: explicit context for printing;
wenzelm
parents: 32187
diff changeset
   167
    SOME th => Drule.implies_intr_list casms
bc341bbe4417 Goal.finish: explicit context for printing;
wenzelm
parents: 32187
diff changeset
   168
      (finish (Syntax.init_pretty_global (Thm.theory_of_thm th)) th)
20250
c3f209752749 prove: proper assumption context, more tactic arguments;
wenzelm
parents: 20228
diff changeset
   169
  | NONE => error "Tactic failed.");
c3f209752749 prove: proper assumption context, more tactic arguments;
wenzelm
parents: 20228
diff changeset
   170
c3f209752749 prove: proper assumption context, more tactic arguments;
wenzelm
parents: 20228
diff changeset
   171
28340
e8597242f649 added promise_result, prove_promise;
wenzelm
parents: 27218
diff changeset
   172
(* prove_common etc. *)
17986
0450847646c3 prove_raw: cterms, explicit asms;
wenzelm
parents: 17983
diff changeset
   173
28340
e8597242f649 added promise_result, prove_promise;
wenzelm
parents: 27218
diff changeset
   174
fun prove_common immediate ctxt xs asms props tac =
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
   175
  let
21516
c2a116a2c4fd ProofContext.init;
wenzelm
parents: 20290
diff changeset
   176
    val thy = ProofContext.theory_of ctxt;
26939
1035c89b4c02 moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents: 26713
diff changeset
   177
    val string_of_term = Syntax.string_of_term ctxt;
20056
0698a403a066 prove/prove_multi: context;
wenzelm
parents: 19862
diff changeset
   178
28355
b9d9360e2440 prove: error with original thread position;
wenzelm
parents: 28341
diff changeset
   179
    val pos = Position.thread_data ();
20250
c3f209752749 prove: proper assumption context, more tactic arguments;
wenzelm
parents: 20228
diff changeset
   180
    fun err msg = cat_error msg
c3f209752749 prove: proper assumption context, more tactic arguments;
wenzelm
parents: 20228
diff changeset
   181
      ("The error(s) above occurred for the goal statement:\n" ^
28355
b9d9360e2440 prove: error with original thread position;
wenzelm
parents: 28341
diff changeset
   182
        string_of_term (Logic.list_implies (asms, Logic.mk_conjunction_list props)) ^
b9d9360e2440 prove: error with original thread position;
wenzelm
parents: 28341
diff changeset
   183
        (case Position.str_of pos of "" => "" | s => "\n" ^ s));
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
   184
20250
c3f209752749 prove: proper assumption context, more tactic arguments;
wenzelm
parents: 20228
diff changeset
   185
    fun cert_safe t = Thm.cterm_of thy (Envir.beta_norm (Term.no_dummy_patterns t))
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
   186
      handle TERM (msg, _) => err msg | TYPE (msg, _, _) => err msg;
20250
c3f209752749 prove: proper assumption context, more tactic arguments;
wenzelm
parents: 20228
diff changeset
   187
    val casms = map cert_safe asms;
c3f209752749 prove: proper assumption context, more tactic arguments;
wenzelm
parents: 20228
diff changeset
   188
    val cprops = map cert_safe props;
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
   189
20250
c3f209752749 prove: proper assumption context, more tactic arguments;
wenzelm
parents: 20228
diff changeset
   190
    val (prems, ctxt') = ctxt
c3f209752749 prove: proper assumption context, more tactic arguments;
wenzelm
parents: 20228
diff changeset
   191
      |> Variable.add_fixes_direct xs
27218
4548c83cd508 prove: full Variable.declare_term, including constraints;
wenzelm
parents: 26939
diff changeset
   192
      |> fold Variable.declare_term (asms @ props)
26713
1c6181def1d0 prove_global: Variable.set_body true, pass context;
wenzelm
parents: 26628
diff changeset
   193
      |> Assumption.add_assumes casms
1c6181def1d0 prove_global: Variable.set_body true, pass context;
wenzelm
parents: 26628
diff changeset
   194
      ||> Variable.set_body true;
28619
89f9dd800a22 prove_common: include all sorts from context into statement, check shyps of result;
wenzelm
parents: 28446
diff changeset
   195
    val sorts = Variable.sorts_of ctxt';
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
   196
28619
89f9dd800a22 prove_common: include all sorts from context into statement, check shyps of result;
wenzelm
parents: 28446
diff changeset
   197
    val stmt = Thm.weaken_sorts sorts (Conjunction.mk_conjunction_balanced cprops);
28340
e8597242f649 added promise_result, prove_promise;
wenzelm
parents: 27218
diff changeset
   198
e8597242f649 added promise_result, prove_promise;
wenzelm
parents: 27218
diff changeset
   199
    fun result () =
e8597242f649 added promise_result, prove_promise;
wenzelm
parents: 27218
diff changeset
   200
      (case SINGLE (tac {prems = prems, context = ctxt'}) (init stmt) of
19774
5fe7731d0836 allow non-trivial schematic goals (via embedded term vars);
wenzelm
parents: 19619
diff changeset
   201
        NONE => err "Tactic failed."
28340
e8597242f649 added promise_result, prove_promise;
wenzelm
parents: 27218
diff changeset
   202
      | SOME st =>
32197
bc341bbe4417 Goal.finish: explicit context for printing;
wenzelm
parents: 32187
diff changeset
   203
          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
   204
            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
   205
            then Thm.check_shyps sorts res
28340
e8597242f649 added promise_result, prove_promise;
wenzelm
parents: 27218
diff changeset
   206
            else err ("Proved a different theorem: " ^ string_of_term (Thm.prop_of res))
e8597242f649 added promise_result, prove_promise;
wenzelm
parents: 27218
diff changeset
   207
          end);
e8597242f649 added promise_result, prove_promise;
wenzelm
parents: 27218
diff changeset
   208
    val res =
29448
34b9652b2f45 added Goal.future_enabled abstraction -- now also checks that this is already
wenzelm
parents: 29435
diff changeset
   209
      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
   210
      then result ()
37186
349e9223c685 explicit markup for forked goals, as indicated by Goal.fork;
wenzelm
parents: 36613
diff changeset
   211
      else future_result ctxt' (fork result) (Thm.term_of stmt);
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
   212
  in
28340
e8597242f649 added promise_result, prove_promise;
wenzelm
parents: 27218
diff changeset
   213
    Conjunction.elim_balanced (length props) res
20290
3f886c176c9b normalized Proof.context/method type aliases;
wenzelm
parents: 20260
diff changeset
   214
    |> map (Assumption.export false ctxt' ctxt)
20056
0698a403a066 prove/prove_multi: context;
wenzelm
parents: 19862
diff changeset
   215
    |> Variable.export ctxt' ctxt
20250
c3f209752749 prove: proper assumption context, more tactic arguments;
wenzelm
parents: 20228
diff changeset
   216
    |> map Drule.zero_var_indexes
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
   217
  end;
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
   218
28341
383f512314b9 prove_multi: immediate;
wenzelm
parents: 28340
diff changeset
   219
val prove_multi = prove_common true;
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
   220
28446
a01de3b3fa2e renamed promise to future, tuned related interfaces;
wenzelm
parents: 28363
diff changeset
   221
fun prove_future ctxt xs asms prop tac = hd (prove_common false ctxt xs asms [prop] tac);
28340
e8597242f649 added promise_result, prove_promise;
wenzelm
parents: 27218
diff changeset
   222
fun prove ctxt xs asms prop tac = hd (prove_common true ctxt xs asms [prop] tac);
20056
0698a403a066 prove/prove_multi: context;
wenzelm
parents: 19862
diff changeset
   223
0698a403a066 prove/prove_multi: context;
wenzelm
parents: 19862
diff changeset
   224
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
   225
  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
   226
09ab79d4e8e1 renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17986
diff changeset
   227
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
   228
21687
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   229
(** goal structure **)
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   230
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   231
(* nested goals *)
18207
9edbeda7e39a tuned norm_hhf_protected;
wenzelm
parents: 18180
diff changeset
   232
19184
3e30297e1300 added extract, retrofit;
wenzelm
parents: 18678
diff changeset
   233
fun extract i n st =
3e30297e1300 added extract, retrofit;
wenzelm
parents: 18678
diff changeset
   234
  (if i < 1 orelse n < 1 orelse i + n - 1 > Thm.nprems_of st then Seq.empty
3e30297e1300 added extract, retrofit;
wenzelm
parents: 18678
diff changeset
   235
   else if n = 1 then Seq.single (Thm.cprem_of st i)
23418
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
   236
   else
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
   237
     Seq.single (Conjunction.mk_conjunction_balanced (map (Thm.cprem_of st) (i upto i + n - 1))))
20260
990dbc007ca6 Thm.adjust_maxidx;
wenzelm
parents: 20250
diff changeset
   238
  |> Seq.map (Thm.adjust_maxidx_cterm ~1 #> init);
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
   239
19221
wenzelm
parents: 19191
diff changeset
   240
fun retrofit i n st' st =
wenzelm
parents: 19191
diff changeset
   241
  (if n = 1 then st
23418
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
   242
   else st |> Drule.with_subgoal i (Conjunction.uncurry_balanced n))
19221
wenzelm
parents: 19191
diff changeset
   243
  |> Thm.compose_no_flatten false (conclude st', Thm.nprems_of st') i;
18207
9edbeda7e39a tuned norm_hhf_protected;
wenzelm
parents: 18180
diff changeset
   244
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
   245
fun SELECT_GOAL tac i st =
19191
56cda3ec2ef8 SELECT_GOAL: fixed trivial case;
wenzelm
parents: 19184
diff changeset
   246
  if Thm.nprems_of st = 1 andalso i = 1 then tac st
19184
3e30297e1300 added extract, retrofit;
wenzelm
parents: 18678
diff changeset
   247
  else Seq.lifts (retrofit i 1) (Seq.maps tac (extract i 1 st)) st;
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
   248
21687
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   249
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   250
(* multiple goals *)
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   251
23418
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
   252
fun precise_conjunction_tac 0 i = eq_assume_tac i
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
   253
  | precise_conjunction_tac 1 i = SUBGOAL (K all_tac) i
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
   254
  | 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
   255
23418
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
   256
val adhoc_conjunction_tac = REPEAT_ALL_NEW
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
   257
  (SUBGOAL (fn (goal, i) =>
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
   258
    if can Logic.dest_conjunction goal then rtac Conjunction.conjunctionI i
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
   259
    else no_tac));
21687
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   260
23418
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
   261
val conjunction_tac = SUBGOAL (fn (goal, i) =>
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
   262
  precise_conjunction_tac (length (Logic.dest_conjunctions goal)) i ORELSE
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
   263
  TRY (adhoc_conjunction_tac i));
21687
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   264
23418
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
   265
val recover_conjunction_tac = PRIMITIVE (fn th =>
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
   266
  Conjunction.uncurry_balanced (Thm.nprems_of th) th);
23414
927203ad4b3a tuned conjunction tactics: slightly smaller proof terms;
wenzelm
parents: 23356
diff changeset
   267
927203ad4b3a tuned conjunction tactics: slightly smaller proof terms;
wenzelm
parents: 23356
diff changeset
   268
fun PRECISE_CONJUNCTS n tac =
927203ad4b3a tuned conjunction tactics: slightly smaller proof terms;
wenzelm
parents: 23356
diff changeset
   269
  SELECT_GOAL (precise_conjunction_tac n 1
927203ad4b3a tuned conjunction tactics: slightly smaller proof terms;
wenzelm
parents: 23356
diff changeset
   270
    THEN tac
23418
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
   271
    THEN recover_conjunction_tac);
23414
927203ad4b3a tuned conjunction tactics: slightly smaller proof terms;
wenzelm
parents: 23356
diff changeset
   272
21687
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   273
fun CONJUNCTS tac =
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   274
  SELECT_GOAL (conjunction_tac 1
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   275
    THEN tac
23418
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
   276
    THEN recover_conjunction_tac);
21687
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   277
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   278
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   279
(* hhf normal form *)
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   280
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   281
val norm_hhf_tac =
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   282
  rtac Drule.asm_rl  (*cheap approximation -- thanks to builtin Logic.flatten_params*)
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   283
  THEN' SUBGOAL (fn (t, i) =>
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   284
    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: 28446
diff changeset
   285
    else MetaSimplifier.rewrite_goal_tac Drule.norm_hhf_eqs i);
21687
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   286
25301
24e027f55f45 removed unused compose_hhf, comp_hhf;
wenzelm
parents: 23536
diff changeset
   287
fun compose_hhf_tac th i st =
24e027f55f45 removed unused compose_hhf, comp_hhf;
wenzelm
parents: 23536
diff changeset
   288
  PRIMSEQ (Thm.bicompose false (false, Drule.lift_all (Thm.cprem_of st i) th, 0) i) st;
21687
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   289
23237
ac9d126456e1 added assume_rule_tac;
wenzelm
parents: 22902
diff changeset
   290
ac9d126456e1 added assume_rule_tac;
wenzelm
parents: 22902
diff changeset
   291
(* non-atomic goal assumptions *)
ac9d126456e1 added assume_rule_tac;
wenzelm
parents: 22902
diff changeset
   292
23356
dbe3731241c3 renamed prove_raw to prove_internal;
wenzelm
parents: 23237
diff changeset
   293
fun non_atomic (Const ("==>", _) $ _ $ _) = true
dbe3731241c3 renamed prove_raw to prove_internal;
wenzelm
parents: 23237
diff changeset
   294
  | non_atomic (Const ("all", _) $ _) = true
dbe3731241c3 renamed prove_raw to prove_internal;
wenzelm
parents: 23237
diff changeset
   295
  | non_atomic _ = false;
dbe3731241c3 renamed prove_raw to prove_internal;
wenzelm
parents: 23237
diff changeset
   296
23237
ac9d126456e1 added assume_rule_tac;
wenzelm
parents: 22902
diff changeset
   297
fun assume_rule_tac ctxt = norm_hhf_tac THEN' CSUBGOAL (fn (goal, i) =>
ac9d126456e1 added assume_rule_tac;
wenzelm
parents: 22902
diff changeset
   298
  let
ac9d126456e1 added assume_rule_tac;
wenzelm
parents: 22902
diff changeset
   299
    val ((_, goal'), ctxt') = Variable.focus goal ctxt;
ac9d126456e1 added assume_rule_tac;
wenzelm
parents: 22902
diff changeset
   300
    val goal'' = Drule.cterm_rule (singleton (Variable.export ctxt' ctxt)) goal';
23356
dbe3731241c3 renamed prove_raw to prove_internal;
wenzelm
parents: 23237
diff changeset
   301
    val Rs = filter (non_atomic o Thm.term_of) (Drule.strip_imp_prems goal'');
23237
ac9d126456e1 added assume_rule_tac;
wenzelm
parents: 22902
diff changeset
   302
    val tacs = Rs |> map (fn R =>
ac9d126456e1 added assume_rule_tac;
wenzelm
parents: 22902
diff changeset
   303
      Tactic.etac (MetaSimplifier.norm_hhf (Thm.trivial R)) THEN_ALL_NEW assume_tac);
ac9d126456e1 added assume_rule_tac;
wenzelm
parents: 22902
diff changeset
   304
  in fold_rev (curry op APPEND') tacs (K no_tac) i end);
ac9d126456e1 added assume_rule_tac;
wenzelm
parents: 22902
diff changeset
   305
32365
9b74d0339c44 added PARALLEL_CHOICE, PARALLEL_GOALS;
wenzelm
parents: 32255
diff changeset
   306
9b74d0339c44 added PARALLEL_CHOICE, PARALLEL_GOALS;
wenzelm
parents: 32255
diff changeset
   307
(* parallel tacticals *)
9b74d0339c44 added PARALLEL_CHOICE, PARALLEL_GOALS;
wenzelm
parents: 32255
diff changeset
   308
9b74d0339c44 added PARALLEL_CHOICE, PARALLEL_GOALS;
wenzelm
parents: 32255
diff changeset
   309
(*parallel choice of single results*)
9b74d0339c44 added PARALLEL_CHOICE, PARALLEL_GOALS;
wenzelm
parents: 32255
diff changeset
   310
fun PARALLEL_CHOICE tacs st =
9b74d0339c44 added PARALLEL_CHOICE, PARALLEL_GOALS;
wenzelm
parents: 32255
diff changeset
   311
  (case Par_List.get_some (fn tac => SINGLE tac st) tacs of
9b74d0339c44 added PARALLEL_CHOICE, PARALLEL_GOALS;
wenzelm
parents: 32255
diff changeset
   312
    NONE => Seq.empty
9b74d0339c44 added PARALLEL_CHOICE, PARALLEL_GOALS;
wenzelm
parents: 32255
diff changeset
   313
  | SOME st' => Seq.single st');
9b74d0339c44 added PARALLEL_CHOICE, PARALLEL_GOALS;
wenzelm
parents: 32255
diff changeset
   314
9b74d0339c44 added PARALLEL_CHOICE, PARALLEL_GOALS;
wenzelm
parents: 32255
diff changeset
   315
(*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
   316
exception FAILED of unit;
32365
9b74d0339c44 added PARALLEL_CHOICE, PARALLEL_GOALS;
wenzelm
parents: 32255
diff changeset
   317
fun PARALLEL_GOALS tac st =
9b74d0339c44 added PARALLEL_CHOICE, PARALLEL_GOALS;
wenzelm
parents: 32255
diff changeset
   318
  let
9b74d0339c44 added PARALLEL_CHOICE, PARALLEL_GOALS;
wenzelm
parents: 32255
diff changeset
   319
    val st0 = Thm.adjust_maxidx_thm ~1 st;
9b74d0339c44 added PARALLEL_CHOICE, PARALLEL_GOALS;
wenzelm
parents: 32255
diff changeset
   320
    val _ = Thm.maxidx_of st0 >= 0 andalso
9b74d0339c44 added PARALLEL_CHOICE, PARALLEL_GOALS;
wenzelm
parents: 32255
diff changeset
   321
      raise THM ("PARALLEL_GOALS: schematic goal state", 0, [st]);
9b74d0339c44 added PARALLEL_CHOICE, PARALLEL_GOALS;
wenzelm
parents: 32255
diff changeset
   322
9b74d0339c44 added PARALLEL_CHOICE, PARALLEL_GOALS;
wenzelm
parents: 32255
diff changeset
   323
    fun try_tac g =
9b74d0339c44 added PARALLEL_CHOICE, PARALLEL_GOALS;
wenzelm
parents: 32255
diff changeset
   324
      (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
   325
        NONE => raise FAILED ()
32365
9b74d0339c44 added PARALLEL_CHOICE, PARALLEL_GOALS;
wenzelm
parents: 32255
diff changeset
   326
      | SOME g' => g');
9b74d0339c44 added PARALLEL_CHOICE, PARALLEL_GOALS;
wenzelm
parents: 32255
diff changeset
   327
9b74d0339c44 added PARALLEL_CHOICE, PARALLEL_GOALS;
wenzelm
parents: 32255
diff changeset
   328
    val goals = Drule.strip_imp_prems (Thm.cprop_of st0);
9b74d0339c44 added PARALLEL_CHOICE, PARALLEL_GOALS;
wenzelm
parents: 32255
diff changeset
   329
    val results = Par_List.map (try_tac o init) goals;
9b74d0339c44 added PARALLEL_CHOICE, PARALLEL_GOALS;
wenzelm
parents: 32255
diff changeset
   330
  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
   331
  handle FAILED () => Seq.empty;
32365
9b74d0339c44 added PARALLEL_CHOICE, PARALLEL_GOALS;
wenzelm
parents: 32255
diff changeset
   332
18207
9edbeda7e39a tuned norm_hhf_protected;
wenzelm
parents: 18180
diff changeset
   333
end;
9edbeda7e39a tuned norm_hhf_protected;
wenzelm
parents: 18180
diff changeset
   334
32365
9b74d0339c44 added PARALLEL_CHOICE, PARALLEL_GOALS;
wenzelm
parents: 32255
diff changeset
   335
structure Basic_Goal: BASIC_GOAL = Goal;
9b74d0339c44 added PARALLEL_CHOICE, PARALLEL_GOALS;
wenzelm
parents: 32255
diff changeset
   336
open Basic_Goal;