src/Pure/goal.ML
author wenzelm
Tue, 19 Jun 2007 23:15:27 +0200
changeset 23418 c195f6f13769
parent 23414 927203ad4b3a
child 23536 60a1672e298e
permissions -rw-r--r--
balanced conjunctions;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/goal.ML
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
     3
    Author:     Makarius and Lawrence C Paulson
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
     4
18139
wenzelm
parents: 18119
diff changeset
     5
Goals in tactical theorem proving.
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
     6
*)
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
     7
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
     8
signature BASIC_GOAL =
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
     9
sig
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
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    13
end;
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    14
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    15
signature GOAL =
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    16
sig
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    17
  include BASIC_GOAL
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    18
  val init: cterm -> thm
18027
09ab79d4e8e1 renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17986
diff changeset
    19
  val protect: thm -> thm
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    20
  val conclude: thm -> thm
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    21
  val finish: thm -> thm
21604
1af327306c8e added norm/close_result (supercede local_standard etc.);
wenzelm
parents: 21579
diff changeset
    22
  val norm_result: thm -> thm
1af327306c8e added norm/close_result (supercede local_standard etc.);
wenzelm
parents: 21579
diff changeset
    23
  val close_result: thm -> thm
23356
dbe3731241c3 renamed prove_raw to prove_internal;
wenzelm
parents: 23237
diff changeset
    24
  val prove_internal: cterm list -> cterm -> (thm list -> tactic) -> thm
20290
3f886c176c9b normalized Proof.context/method type aliases;
wenzelm
parents: 20260
diff changeset
    25
  val prove_multi: Proof.context -> string list -> term list -> term list ->
3f886c176c9b normalized Proof.context/method type aliases;
wenzelm
parents: 20260
diff changeset
    26
    ({prems: thm list, context: Proof.context} -> tactic) -> thm list
3f886c176c9b normalized Proof.context/method type aliases;
wenzelm
parents: 20260
diff changeset
    27
  val prove: Proof.context -> string list -> term list -> term ->
3f886c176c9b normalized Proof.context/method type aliases;
wenzelm
parents: 20260
diff changeset
    28
    ({prems: thm list, context: Proof.context} -> tactic) -> thm
20056
0698a403a066 prove/prove_multi: context;
wenzelm
parents: 19862
diff changeset
    29
  val prove_global: theory -> string list -> term list -> term -> (thm list -> tactic) -> thm
19184
3e30297e1300 added extract, retrofit;
wenzelm
parents: 18678
diff changeset
    30
  val extract: int -> int -> thm -> thm Seq.seq
3e30297e1300 added extract, retrofit;
wenzelm
parents: 18678
diff changeset
    31
  val retrofit: int -> int -> thm -> thm -> thm Seq.seq
23418
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
    32
  val conjunction_tac: int -> tactic
23414
927203ad4b3a tuned conjunction tactics: slightly smaller proof terms;
wenzelm
parents: 23356
diff changeset
    33
  val precise_conjunction_tac: int -> int -> tactic
23418
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
    34
  val recover_conjunction_tac: tactic
21687
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
    35
  val asm_rewrite_goal_tac: bool * bool * bool -> (simpset -> tactic) -> simpset -> int -> tactic
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
    36
  val rewrite_goal_tac: thm list -> int -> tactic
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
    37
  val norm_hhf_tac: int -> tactic
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
    38
  val compose_hhf: thm -> int -> thm -> thm Seq.seq
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
    39
  val compose_hhf_tac: thm -> int -> tactic
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
    40
  val comp_hhf: thm -> thm -> thm
23237
ac9d126456e1 added assume_rule_tac;
wenzelm
parents: 22902
diff changeset
    41
  val assume_rule_tac: Proof.context -> int -> tactic
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    42
end;
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    43
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    44
structure Goal: GOAL =
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    45
struct
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    46
18027
09ab79d4e8e1 renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17986
diff changeset
    47
(** goals **)
09ab79d4e8e1 renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17986
diff changeset
    48
09ab79d4e8e1 renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17986
diff changeset
    49
(*
09ab79d4e8e1 renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17986
diff changeset
    50
  -------- (init)
09ab79d4e8e1 renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17986
diff changeset
    51
  C ==> #C
09ab79d4e8e1 renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17986
diff changeset
    52
*)
20290
3f886c176c9b normalized Proof.context/method type aliases;
wenzelm
parents: 20260
diff changeset
    53
val init =
22902
ac833b4bb7ee moved some Drule operations to Thm (see more_thm.ML);
wenzelm
parents: 21687
diff changeset
    54
  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
    55
  in fn C => Thm.instantiate ([], [(A, C)]) Drule.protectI end;
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    56
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    57
(*
18119
63901a9b99dc export compose_hhf;
wenzelm
parents: 18027
diff changeset
    58
   C
63901a9b99dc export compose_hhf;
wenzelm
parents: 18027
diff changeset
    59
  --- (protect)
63901a9b99dc export compose_hhf;
wenzelm
parents: 18027
diff changeset
    60
  #C
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    61
*)
21579
abd2b4386a63 COMP_INCR;
wenzelm
parents: 21516
diff changeset
    62
fun protect th = th COMP_INCR Drule.protectI;
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    63
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    64
(*
18027
09ab79d4e8e1 renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17986
diff changeset
    65
  A ==> ... ==> #C
09ab79d4e8e1 renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17986
diff changeset
    66
  ---------------- (conclude)
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    67
  A ==> ... ==> C
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    68
*)
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    69
fun conclude th =
18497
7569674d7bb1 Thm.compose_no_flatten;
wenzelm
parents: 18484
diff changeset
    70
  (case SINGLE (Thm.compose_no_flatten false (th, Thm.nprems_of th) 1)
18027
09ab79d4e8e1 renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17986
diff changeset
    71
      (Drule.incr_indexes th Drule.protectD) of
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    72
    SOME th' => th'
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    73
  | NONE => raise THM ("Failed to conclude goal", 0, [th]));
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    74
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    75
(*
18027
09ab79d4e8e1 renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17986
diff changeset
    76
  #C
09ab79d4e8e1 renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17986
diff changeset
    77
  --- (finish)
09ab79d4e8e1 renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17986
diff changeset
    78
   C
17983
wenzelm
parents: 17980
diff changeset
    79
*)
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    80
fun finish th =
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    81
  (case Thm.nprems_of th of
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    82
    0 => conclude th
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    83
  | n => raise THM ("Proof failed.\n" ^
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    84
      Pretty.string_of (Pretty.chunks (Display.pretty_goals n th)) ^
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    85
      ("\n" ^ string_of_int n ^ " unsolved goal(s)!"), 0, [th]));
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    86
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
    87
18027
09ab79d4e8e1 renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17986
diff changeset
    88
09ab79d4e8e1 renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17986
diff changeset
    89
(** results **)
09ab79d4e8e1 renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17986
diff changeset
    90
21604
1af327306c8e added norm/close_result (supercede local_standard etc.);
wenzelm
parents: 21579
diff changeset
    91
(* normal form *)
1af327306c8e added norm/close_result (supercede local_standard etc.);
wenzelm
parents: 21579
diff changeset
    92
1af327306c8e added norm/close_result (supercede local_standard etc.);
wenzelm
parents: 21579
diff changeset
    93
val norm_result =
1af327306c8e added norm/close_result (supercede local_standard etc.);
wenzelm
parents: 21579
diff changeset
    94
  Drule.flexflex_unique
1af327306c8e added norm/close_result (supercede local_standard etc.);
wenzelm
parents: 21579
diff changeset
    95
  #> MetaSimplifier.norm_hhf_protect
1af327306c8e added norm/close_result (supercede local_standard etc.);
wenzelm
parents: 21579
diff changeset
    96
  #> Thm.strip_shyps
1af327306c8e added norm/close_result (supercede local_standard etc.);
wenzelm
parents: 21579
diff changeset
    97
  #> Drule.zero_var_indexes;
1af327306c8e added norm/close_result (supercede local_standard etc.);
wenzelm
parents: 21579
diff changeset
    98
1af327306c8e added norm/close_result (supercede local_standard etc.);
wenzelm
parents: 21579
diff changeset
    99
val close_result =
1af327306c8e added norm/close_result (supercede local_standard etc.);
wenzelm
parents: 21579
diff changeset
   100
  Thm.compress
1af327306c8e added norm/close_result (supercede local_standard etc.);
wenzelm
parents: 21579
diff changeset
   101
  #> Drule.close_derivation;
1af327306c8e added norm/close_result (supercede local_standard etc.);
wenzelm
parents: 21579
diff changeset
   102
1af327306c8e added norm/close_result (supercede local_standard etc.);
wenzelm
parents: 21579
diff changeset
   103
18027
09ab79d4e8e1 renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17986
diff changeset
   104
09ab79d4e8e1 renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17986
diff changeset
   105
(** tactical theorem proving **)
09ab79d4e8e1 renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17986
diff changeset
   106
23356
dbe3731241c3 renamed prove_raw to prove_internal;
wenzelm
parents: 23237
diff changeset
   107
(* prove_internal -- minimal checks, no normalization of result! *)
20250
c3f209752749 prove: proper assumption context, more tactic arguments;
wenzelm
parents: 20228
diff changeset
   108
23356
dbe3731241c3 renamed prove_raw to prove_internal;
wenzelm
parents: 23237
diff changeset
   109
fun prove_internal casms cprop tac =
20250
c3f209752749 prove: proper assumption context, more tactic arguments;
wenzelm
parents: 20228
diff changeset
   110
  (case SINGLE (tac (map Assumption.assume casms)) (init cprop) of
c3f209752749 prove: proper assumption context, more tactic arguments;
wenzelm
parents: 20228
diff changeset
   111
    SOME th => Drule.implies_intr_list casms (finish th)
c3f209752749 prove: proper assumption context, more tactic arguments;
wenzelm
parents: 20228
diff changeset
   112
  | NONE => error "Tactic failed.");
c3f209752749 prove: proper assumption context, more tactic arguments;
wenzelm
parents: 20228
diff changeset
   113
c3f209752749 prove: proper assumption context, more tactic arguments;
wenzelm
parents: 20228
diff changeset
   114
18119
63901a9b99dc export compose_hhf;
wenzelm
parents: 18027
diff changeset
   115
(* prove_multi *)
17986
0450847646c3 prove_raw: cterms, explicit asms;
wenzelm
parents: 17983
diff changeset
   116
20056
0698a403a066 prove/prove_multi: context;
wenzelm
parents: 19862
diff changeset
   117
fun prove_multi ctxt xs asms props tac =
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
   118
  let
21516
c2a116a2c4fd ProofContext.init;
wenzelm
parents: 20290
diff changeset
   119
    val thy = ProofContext.theory_of ctxt;
20056
0698a403a066 prove/prove_multi: context;
wenzelm
parents: 19862
diff changeset
   120
    val string_of_term = Sign.string_of_term thy;
0698a403a066 prove/prove_multi: context;
wenzelm
parents: 19862
diff changeset
   121
20250
c3f209752749 prove: proper assumption context, more tactic arguments;
wenzelm
parents: 20228
diff changeset
   122
    fun err msg = cat_error msg
c3f209752749 prove: proper assumption context, more tactic arguments;
wenzelm
parents: 20228
diff changeset
   123
      ("The error(s) above occurred for the goal statement:\n" ^
c3f209752749 prove: proper assumption context, more tactic arguments;
wenzelm
parents: 20228
diff changeset
   124
        string_of_term (Logic.list_implies (asms, Logic.mk_conjunction_list props)));
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
   125
20250
c3f209752749 prove: proper assumption context, more tactic arguments;
wenzelm
parents: 20228
diff changeset
   126
    fun cert_safe t = Thm.cterm_of thy (Envir.beta_norm (Term.no_dummy_patterns t))
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
   127
      handle TERM (msg, _) => err msg | TYPE (msg, _, _) => err msg;
20250
c3f209752749 prove: proper assumption context, more tactic arguments;
wenzelm
parents: 20228
diff changeset
   128
    val casms = map cert_safe asms;
c3f209752749 prove: proper assumption context, more tactic arguments;
wenzelm
parents: 20228
diff changeset
   129
    val cprops = map cert_safe props;
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
   130
20250
c3f209752749 prove: proper assumption context, more tactic arguments;
wenzelm
parents: 20228
diff changeset
   131
    val (prems, ctxt') = ctxt
c3f209752749 prove: proper assumption context, more tactic arguments;
wenzelm
parents: 20228
diff changeset
   132
      |> Variable.add_fixes_direct xs
c3f209752749 prove: proper assumption context, more tactic arguments;
wenzelm
parents: 20228
diff changeset
   133
      |> fold Variable.declare_internal (asms @ props)
c3f209752749 prove: proper assumption context, more tactic arguments;
wenzelm
parents: 20228
diff changeset
   134
      |> Assumption.add_assumes casms;
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
   135
23418
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
   136
    val goal = init (Conjunction.mk_conjunction_balanced cprops);
19774
5fe7731d0836 allow non-trivial schematic goals (via embedded term vars);
wenzelm
parents: 19619
diff changeset
   137
    val res =
20250
c3f209752749 prove: proper assumption context, more tactic arguments;
wenzelm
parents: 20228
diff changeset
   138
      (case SINGLE (tac {prems = prems, context = ctxt'}) goal of
19774
5fe7731d0836 allow non-trivial schematic goals (via embedded term vars);
wenzelm
parents: 19619
diff changeset
   139
        NONE => err "Tactic failed."
5fe7731d0836 allow non-trivial schematic goals (via embedded term vars);
wenzelm
parents: 19619
diff changeset
   140
      | SOME res => res);
23418
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
   141
    val results = Conjunction.elim_balanced (length props) (finish res)
20250
c3f209752749 prove: proper assumption context, more tactic arguments;
wenzelm
parents: 20228
diff changeset
   142
      handle THM (msg, _, _) => err msg;
c3f209752749 prove: proper assumption context, more tactic arguments;
wenzelm
parents: 20228
diff changeset
   143
    val _ = Unify.matches_list thy (map Thm.term_of cprops) (map Thm.prop_of results)
20056
0698a403a066 prove/prove_multi: context;
wenzelm
parents: 19862
diff changeset
   144
      orelse err ("Proved a different theorem: " ^ string_of_term (Thm.prop_of res));
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
   145
  in
20056
0698a403a066 prove/prove_multi: context;
wenzelm
parents: 19862
diff changeset
   146
    results
20290
3f886c176c9b normalized Proof.context/method type aliases;
wenzelm
parents: 20260
diff changeset
   147
    |> map (Assumption.export false ctxt' ctxt)
20056
0698a403a066 prove/prove_multi: context;
wenzelm
parents: 19862
diff changeset
   148
    |> Variable.export ctxt' ctxt
20250
c3f209752749 prove: proper assumption context, more tactic arguments;
wenzelm
parents: 20228
diff changeset
   149
    |> map Drule.zero_var_indexes
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
   150
  end;
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
   151
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
   152
18119
63901a9b99dc export compose_hhf;
wenzelm
parents: 18027
diff changeset
   153
(* prove *)
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
   154
20056
0698a403a066 prove/prove_multi: context;
wenzelm
parents: 19862
diff changeset
   155
fun prove ctxt xs asms prop tac = hd (prove_multi ctxt xs asms [prop] tac);
0698a403a066 prove/prove_multi: context;
wenzelm
parents: 19862
diff changeset
   156
0698a403a066 prove/prove_multi: context;
wenzelm
parents: 19862
diff changeset
   157
fun prove_global thy xs asms prop tac =
21516
c2a116a2c4fd ProofContext.init;
wenzelm
parents: 20290
diff changeset
   158
  Drule.standard (prove (ProofContext.init thy) xs asms prop (fn {prems, ...} => tac prems));
18027
09ab79d4e8e1 renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17986
diff changeset
   159
09ab79d4e8e1 renamed Goal constant to prop, more general protect/unprotect interfaces;
wenzelm
parents: 17986
diff changeset
   160
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
   161
21687
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   162
(** goal structure **)
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   163
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   164
(* nested goals *)
18207
9edbeda7e39a tuned norm_hhf_protected;
wenzelm
parents: 18180
diff changeset
   165
19184
3e30297e1300 added extract, retrofit;
wenzelm
parents: 18678
diff changeset
   166
fun extract i n st =
3e30297e1300 added extract, retrofit;
wenzelm
parents: 18678
diff changeset
   167
  (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
   168
   else if n = 1 then Seq.single (Thm.cprem_of st i)
23418
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
   169
   else
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
   170
     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
   171
  |> Seq.map (Thm.adjust_maxidx_cterm ~1 #> init);
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
   172
19221
wenzelm
parents: 19191
diff changeset
   173
fun retrofit i n st' st =
wenzelm
parents: 19191
diff changeset
   174
  (if n = 1 then st
23418
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
   175
   else st |> Drule.with_subgoal i (Conjunction.uncurry_balanced n))
19221
wenzelm
parents: 19191
diff changeset
   176
  |> Thm.compose_no_flatten false (conclude st', Thm.nprems_of st') i;
18207
9edbeda7e39a tuned norm_hhf_protected;
wenzelm
parents: 18180
diff changeset
   177
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
   178
fun SELECT_GOAL tac i st =
19191
56cda3ec2ef8 SELECT_GOAL: fixed trivial case;
wenzelm
parents: 19184
diff changeset
   179
  if Thm.nprems_of st = 1 andalso i = 1 then tac st
19184
3e30297e1300 added extract, retrofit;
wenzelm
parents: 18678
diff changeset
   180
  else Seq.lifts (retrofit i 1) (Seq.maps tac (extract i 1 st)) st;
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
   181
21687
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   182
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   183
(* multiple goals *)
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   184
23418
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
   185
fun precise_conjunction_tac 0 i = eq_assume_tac i
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
   186
  | precise_conjunction_tac 1 i = SUBGOAL (K all_tac) i
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
   187
  | 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
   188
23418
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
   189
val adhoc_conjunction_tac = REPEAT_ALL_NEW
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
   190
  (SUBGOAL (fn (goal, i) =>
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
   191
    if can Logic.dest_conjunction goal then rtac Conjunction.conjunctionI i
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
   192
    else no_tac));
21687
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   193
23418
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
   194
val conjunction_tac = SUBGOAL (fn (goal, i) =>
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
   195
  precise_conjunction_tac (length (Logic.dest_conjunctions goal)) i ORELSE
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
   196
  TRY (adhoc_conjunction_tac i));
21687
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   197
23418
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
   198
val recover_conjunction_tac = PRIMITIVE (fn th =>
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
   199
  Conjunction.uncurry_balanced (Thm.nprems_of th) th);
23414
927203ad4b3a tuned conjunction tactics: slightly smaller proof terms;
wenzelm
parents: 23356
diff changeset
   200
927203ad4b3a tuned conjunction tactics: slightly smaller proof terms;
wenzelm
parents: 23356
diff changeset
   201
fun PRECISE_CONJUNCTS n tac =
927203ad4b3a tuned conjunction tactics: slightly smaller proof terms;
wenzelm
parents: 23356
diff changeset
   202
  SELECT_GOAL (precise_conjunction_tac n 1
927203ad4b3a tuned conjunction tactics: slightly smaller proof terms;
wenzelm
parents: 23356
diff changeset
   203
    THEN tac
23418
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
   204
    THEN recover_conjunction_tac);
23414
927203ad4b3a tuned conjunction tactics: slightly smaller proof terms;
wenzelm
parents: 23356
diff changeset
   205
21687
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   206
fun CONJUNCTS tac =
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   207
  SELECT_GOAL (conjunction_tac 1
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   208
    THEN tac
23418
c195f6f13769 balanced conjunctions;
wenzelm
parents: 23414
diff changeset
   209
    THEN recover_conjunction_tac);
21687
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   210
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   211
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   212
(* rewriting *)
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   213
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   214
(*Rewrite subgoal i only.  SELECT_GOAL avoids inefficiencies in goals_conv.*)
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   215
fun asm_rewrite_goal_tac mode prover_tac ss =
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   216
  SELECT_GOAL
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   217
    (PRIMITIVE (MetaSimplifier.rewrite_goal_rule mode (SINGLE o prover_tac) ss 1));
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   218
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   219
fun rewrite_goal_tac rews =
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   220
  let val ss = MetaSimplifier.empty_ss addsimps rews in
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   221
    fn i => fn st => asm_rewrite_goal_tac (true, false, false) (K no_tac)
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   222
      (MetaSimplifier.theory_context (Thm.theory_of_thm st) ss) i st
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   223
  end;
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   224
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   225
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   226
(* hhf normal form *)
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   227
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   228
val norm_hhf_tac =
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   229
  rtac Drule.asm_rl  (*cheap approximation -- thanks to builtin Logic.flatten_params*)
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   230
  THEN' SUBGOAL (fn (t, i) =>
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   231
    if Drule.is_norm_hhf t then all_tac
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   232
    else rewrite_goal_tac [Drule.norm_hhf_eq] i);
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   233
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   234
fun compose_hhf tha i thb =
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   235
  Thm.bicompose false (false, Drule.lift_all (Thm.cprem_of thb i) tha, 0) i thb;
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   236
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   237
fun compose_hhf_tac th i = PRIMSEQ (compose_hhf th i);
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   238
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   239
fun comp_hhf tha thb =
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   240
  (case Seq.chop 2 (compose_hhf tha 1 thb) of
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   241
    ([th], _) => th
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   242
  | ([], _) => raise THM ("comp_hhf: no unifiers", 1, [tha, thb])
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   243
  | _  => raise THM ("comp_hhf: multiple unifiers", 1, [tha, thb]));
f689f729afab reorganized structure Goal vs. Tactic;
wenzelm
parents: 21604
diff changeset
   244
23237
ac9d126456e1 added assume_rule_tac;
wenzelm
parents: 22902
diff changeset
   245
ac9d126456e1 added assume_rule_tac;
wenzelm
parents: 22902
diff changeset
   246
(* non-atomic goal assumptions *)
ac9d126456e1 added assume_rule_tac;
wenzelm
parents: 22902
diff changeset
   247
23356
dbe3731241c3 renamed prove_raw to prove_internal;
wenzelm
parents: 23237
diff changeset
   248
fun non_atomic (Const ("==>", _) $ _ $ _) = true
dbe3731241c3 renamed prove_raw to prove_internal;
wenzelm
parents: 23237
diff changeset
   249
  | non_atomic (Const ("all", _) $ _) = true
dbe3731241c3 renamed prove_raw to prove_internal;
wenzelm
parents: 23237
diff changeset
   250
  | non_atomic _ = false;
dbe3731241c3 renamed prove_raw to prove_internal;
wenzelm
parents: 23237
diff changeset
   251
23237
ac9d126456e1 added assume_rule_tac;
wenzelm
parents: 22902
diff changeset
   252
fun assume_rule_tac ctxt = norm_hhf_tac THEN' CSUBGOAL (fn (goal, i) =>
ac9d126456e1 added assume_rule_tac;
wenzelm
parents: 22902
diff changeset
   253
  let
ac9d126456e1 added assume_rule_tac;
wenzelm
parents: 22902
diff changeset
   254
    val ((_, goal'), ctxt') = Variable.focus goal ctxt;
ac9d126456e1 added assume_rule_tac;
wenzelm
parents: 22902
diff changeset
   255
    val goal'' = Drule.cterm_rule (singleton (Variable.export ctxt' ctxt)) goal';
23356
dbe3731241c3 renamed prove_raw to prove_internal;
wenzelm
parents: 23237
diff changeset
   256
    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
   257
    val tacs = Rs |> map (fn R =>
ac9d126456e1 added assume_rule_tac;
wenzelm
parents: 22902
diff changeset
   258
      Tactic.etac (MetaSimplifier.norm_hhf (Thm.trivial R)) THEN_ALL_NEW assume_tac);
ac9d126456e1 added assume_rule_tac;
wenzelm
parents: 22902
diff changeset
   259
  in fold_rev (curry op APPEND') tacs (K no_tac) i end);
ac9d126456e1 added assume_rule_tac;
wenzelm
parents: 22902
diff changeset
   260
18207
9edbeda7e39a tuned norm_hhf_protected;
wenzelm
parents: 18180
diff changeset
   261
end;
9edbeda7e39a tuned norm_hhf_protected;
wenzelm
parents: 18180
diff changeset
   262
17980
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
   263
structure BasicGoal: BASIC_GOAL = Goal;
788836292b1a Internal goals.
wenzelm
parents:
diff changeset
   264
open BasicGoal;