src/Pure/subgoal.ML
author wenzelm
Wed, 01 Jul 2015 22:37:49 +0200
changeset 60627 5d13babbb3f6
parent 60626 9eefb9f39021
child 60628 5ff15d0dd7fa
permissions -rw-r--r--
split multi-goals as usual (outermost Pure.conjunction only);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
20210
8fe4ae4360eb Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/subgoal.ML
8fe4ae4360eb Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff changeset
     2
    Author:     Makarius
60623
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
     3
    Author:     Daniel Matichuk, NICTA/UNSW
20210
8fe4ae4360eb Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff changeset
     4
32281
750101435f60 focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
wenzelm
parents: 32213
diff changeset
     5
Tactical operations with explicit subgoal focus, based on canonical
750101435f60 focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
wenzelm
parents: 32213
diff changeset
     6
proof decomposition.  The "visible" part of the text within the
750101435f60 focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
wenzelm
parents: 32213
diff changeset
     7
context is fixed, the remaining goal may be schematic.
60623
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
     8
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
     9
Isar subgoal command for proof structure within unstructured proof
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
    10
scripts.
20210
8fe4ae4360eb Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff changeset
    11
*)
8fe4ae4360eb Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff changeset
    12
8fe4ae4360eb Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff changeset
    13
signature SUBGOAL =
8fe4ae4360eb Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff changeset
    14
sig
32200
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
    15
  type focus = {context: Proof.context, params: (string * cterm) list, prems: thm list,
32281
750101435f60 focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
wenzelm
parents: 32213
diff changeset
    16
    asms: cterm list, concl: cterm, schematics: (ctyp * ctyp) list * (cterm * cterm) list}
750101435f60 focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
wenzelm
parents: 32213
diff changeset
    17
  val focus_params: Proof.context -> int -> thm -> focus * thm
60626
9eefb9f39021 clarified prems: full subgoal is imported in any case, to avoid remaining schematic variables;
wenzelm
parents: 60625
diff changeset
    18
  val focus_params_fixed: Proof.context -> int -> thm -> focus * thm
32281
750101435f60 focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
wenzelm
parents: 32213
diff changeset
    19
  val focus_prems: Proof.context -> int -> thm -> focus * thm
32200
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
    20
  val focus: Proof.context -> int -> thm -> focus * thm
32210
a5e9d9f3e5e1 retrofit: actually handle schematic variables -- need to export into original context;
wenzelm
parents: 32200
diff changeset
    21
  val retrofit: Proof.context -> Proof.context -> (string * cterm) list -> cterm list ->
32200
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
    22
    int -> thm -> thm -> thm Seq.seq
32281
750101435f60 focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
wenzelm
parents: 32213
diff changeset
    23
  val FOCUS_PARAMS: (focus -> tactic) -> Proof.context -> int -> tactic
60626
9eefb9f39021 clarified prems: full subgoal is imported in any case, to avoid remaining schematic variables;
wenzelm
parents: 60625
diff changeset
    24
  val FOCUS_PARAMS_FIXED: (focus -> tactic) -> Proof.context -> int -> tactic
32281
750101435f60 focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
wenzelm
parents: 32213
diff changeset
    25
  val FOCUS_PREMS: (focus -> tactic) -> Proof.context -> int -> tactic
32200
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
    26
  val FOCUS: (focus -> tactic) -> Proof.context -> int -> tactic
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
    27
  val SUBPROOF: (focus -> tactic) -> Proof.context -> int -> tactic
60623
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
    28
  val subgoal: Attrib.binding -> Attrib.binding option -> (binding * mixfix) list ->
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
    29
    Proof.state -> focus * Proof.state
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
    30
  val subgoal_cmd: Attrib.binding -> Attrib.binding option -> (binding * mixfix) list ->
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
    31
    Proof.state -> focus * Proof.state
20210
8fe4ae4360eb Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff changeset
    32
end;
8fe4ae4360eb Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff changeset
    33
8fe4ae4360eb Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff changeset
    34
structure Subgoal: SUBGOAL =
8fe4ae4360eb Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff changeset
    35
struct
8fe4ae4360eb Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff changeset
    36
32200
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
    37
(* focus *)
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
    38
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
    39
type focus = {context: Proof.context, params: (string * cterm) list, prems: thm list,
32281
750101435f60 focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
wenzelm
parents: 32213
diff changeset
    40
  asms: cterm list, concl: cterm, schematics: (ctyp * ctyp) list * (cterm * cterm) list};
20210
8fe4ae4360eb Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff changeset
    41
32281
750101435f60 focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
wenzelm
parents: 32213
diff changeset
    42
fun gen_focus (do_prems, do_concl) ctxt i raw_st =
20210
8fe4ae4360eb Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff changeset
    43
  let
60367
e201bd8973db clarified context;
wenzelm
parents: 59621
diff changeset
    44
    val st = raw_st
e201bd8973db clarified context;
wenzelm
parents: 59621
diff changeset
    45
      |> Thm.transfer (Proof_Context.theory_of ctxt)
e201bd8973db clarified context;
wenzelm
parents: 59621
diff changeset
    46
      |> Simplifier.norm_hhf_protect ctxt;
32281
750101435f60 focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
wenzelm
parents: 32213
diff changeset
    47
    val ((schematic_types, [st']), ctxt1) = Variable.importT [st] ctxt;
42495
1af81b70cf09 clarified Variable.focus vs. Variable.focus_cterm -- eliminated clone;
wenzelm
parents: 42360
diff changeset
    48
    val ((params, goal), ctxt2) = Variable.focus_cterm (Thm.cprem_of st' i) ctxt1;
32281
750101435f60 focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
wenzelm
parents: 32213
diff changeset
    49
32213
f1b166317ae2 added focus_params/FOCUS_PARAMS, which focus on the parameter prefix only;
wenzelm
parents: 32210
diff changeset
    50
    val (asms, concl) =
32281
750101435f60 focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
wenzelm
parents: 32213
diff changeset
    51
      if do_prems then (Drule.strip_imp_prems goal, Drule.strip_imp_concl goal)
750101435f60 focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
wenzelm
parents: 32213
diff changeset
    52
      else ([], goal);
750101435f60 focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
wenzelm
parents: 32213
diff changeset
    53
    val text = asms @ (if do_concl then [concl] else []);
750101435f60 focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
wenzelm
parents: 32213
diff changeset
    54
60367
e201bd8973db clarified context;
wenzelm
parents: 59621
diff changeset
    55
    val (inst, ctxt3) = Variable.import_inst true (map Thm.term_of text) ctxt2;
e201bd8973db clarified context;
wenzelm
parents: 59621
diff changeset
    56
    val (_, schematic_terms) = Thm.certify_inst ctxt3 inst;
32281
750101435f60 focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
wenzelm
parents: 32213
diff changeset
    57
750101435f60 focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
wenzelm
parents: 32213
diff changeset
    58
    val schematics = (schematic_types, schematic_terms);
750101435f60 focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
wenzelm
parents: 32213
diff changeset
    59
    val asms' = map (Thm.instantiate_cterm schematics) asms;
750101435f60 focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
wenzelm
parents: 32213
diff changeset
    60
    val concl' = Thm.instantiate_cterm schematics concl;
750101435f60 focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
wenzelm
parents: 32213
diff changeset
    61
    val (prems, context) = Assumption.add_assumes asms' ctxt3;
20210
8fe4ae4360eb Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff changeset
    62
  in
32200
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
    63
    ({context = context, params = params, prems = prems,
32281
750101435f60 focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
wenzelm
parents: 32213
diff changeset
    64
      asms = asms', concl = concl', schematics = schematics}, Goal.init concl')
20210
8fe4ae4360eb Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff changeset
    65
  end;
8fe4ae4360eb Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff changeset
    66
32281
750101435f60 focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
wenzelm
parents: 32213
diff changeset
    67
val focus_params = gen_focus (false, false);
60626
9eefb9f39021 clarified prems: full subgoal is imported in any case, to avoid remaining schematic variables;
wenzelm
parents: 60625
diff changeset
    68
val focus_params_fixed = gen_focus (false, true);
32281
750101435f60 focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
wenzelm
parents: 32213
diff changeset
    69
val focus_prems = gen_focus (true, false);
750101435f60 focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
wenzelm
parents: 32213
diff changeset
    70
val focus = gen_focus (true, true);
32213
f1b166317ae2 added focus_params/FOCUS_PARAMS, which focus on the parameter prefix only;
wenzelm
parents: 32210
diff changeset
    71
32200
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
    72
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
    73
(* lift and retrofit *)
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
    74
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
    75
(*
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
    76
     B [?'b, ?y]
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
    77
  ----------------
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
    78
  B ['b, y params]
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
    79
*)
32284
d8ee8a956f19 retrofit: more precise handling of locally introduced schematic variables (avoid zero_var_indexes);
wenzelm
parents: 32283
diff changeset
    80
fun lift_import idx params th ctxt =
32200
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
    81
  let
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
    82
    val ((_, [th']), ctxt') = Variable.importT [th] ctxt;
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
    83
59586
ddf6deaadfe8 clarified signature;
wenzelm
parents: 59573
diff changeset
    84
    val Ts = map Thm.typ_of_cterm params;
32200
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
    85
    val ts = map Thm.term_of params;
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
    86
32281
750101435f60 focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
wenzelm
parents: 32213
diff changeset
    87
    val prop = Thm.full_prop_of th';
750101435f60 focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
wenzelm
parents: 32213
diff changeset
    88
    val concl_vars = Term.add_vars (Logic.strip_imp_concl prop) [];
750101435f60 focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
wenzelm
parents: 32213
diff changeset
    89
    val vars = rev (Term.add_vars prop []);
32200
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
    90
    val (ys, ctxt'') = Variable.variant_fixes (map (Name.clean o #1 o #1) vars) ctxt';
32284
d8ee8a956f19 retrofit: more precise handling of locally introduced schematic variables (avoid zero_var_indexes);
wenzelm
parents: 32283
diff changeset
    91
d8ee8a956f19 retrofit: more precise handling of locally introduced schematic variables (avoid zero_var_indexes);
wenzelm
parents: 32283
diff changeset
    92
    fun var_inst v y =
d8ee8a956f19 retrofit: more precise handling of locally introduced schematic variables (avoid zero_var_indexes);
wenzelm
parents: 32283
diff changeset
    93
      let
d8ee8a956f19 retrofit: more precise handling of locally introduced schematic variables (avoid zero_var_indexes);
wenzelm
parents: 32283
diff changeset
    94
        val ((x, i), T) = v;
d8ee8a956f19 retrofit: more precise handling of locally introduced schematic variables (avoid zero_var_indexes);
wenzelm
parents: 32283
diff changeset
    95
        val (U, args) =
d8ee8a956f19 retrofit: more precise handling of locally introduced schematic variables (avoid zero_var_indexes);
wenzelm
parents: 32283
diff changeset
    96
          if member (op =) concl_vars v then (T, [])
d8ee8a956f19 retrofit: more precise handling of locally introduced schematic variables (avoid zero_var_indexes);
wenzelm
parents: 32283
diff changeset
    97
          else (Ts ---> T, ts);
d8ee8a956f19 retrofit: more precise handling of locally introduced schematic variables (avoid zero_var_indexes);
wenzelm
parents: 32283
diff changeset
    98
        val u = Free (y, U);
d8ee8a956f19 retrofit: more precise handling of locally introduced schematic variables (avoid zero_var_indexes);
wenzelm
parents: 32283
diff changeset
    99
        in ((Var v, list_comb (u, args)), (u, Var ((x, i + idx), U))) end;
59616
eb59c6968219 tuned -- more explicit use of context;
wenzelm
parents: 59586
diff changeset
   100
    val (inst1, inst2) =
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59616
diff changeset
   101
      split_list (map (apply2 (apply2 (Thm.cterm_of ctxt))) (map2 var_inst vars ys));
32284
d8ee8a956f19 retrofit: more precise handling of locally introduced schematic variables (avoid zero_var_indexes);
wenzelm
parents: 32283
diff changeset
   102
d8ee8a956f19 retrofit: more precise handling of locally introduced schematic variables (avoid zero_var_indexes);
wenzelm
parents: 32283
diff changeset
   103
    val th'' = Thm.instantiate ([], inst1) th';
d8ee8a956f19 retrofit: more precise handling of locally introduced schematic variables (avoid zero_var_indexes);
wenzelm
parents: 32283
diff changeset
   104
  in ((inst2, th''), ctxt'') end;
32200
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
   105
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
   106
(*
32281
750101435f60 focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
wenzelm
parents: 32213
diff changeset
   107
       [x, A x]
32200
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
   108
          :
32281
750101435f60 focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
wenzelm
parents: 32213
diff changeset
   109
       B x ==> C
32200
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
   110
  ------------------
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
   111
  [!!x. A x ==> B x]
32281
750101435f60 focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
wenzelm
parents: 32213
diff changeset
   112
          :
750101435f60 focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
wenzelm
parents: 32213
diff changeset
   113
          C
32200
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
   114
*)
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
   115
fun lift_subgoals params asms th =
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
   116
  let
32329
1527ff8c2dfb SUBPROOF: recovered Goal.check_finished;
wenzelm
parents: 32284
diff changeset
   117
    fun lift ct = fold_rev Thm.all_name params (Drule.list_implies (asms, ct));
32200
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
   118
    val unlift =
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
   119
      fold (Thm.elim_implies o Thm.assume) asms o
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
   120
      Drule.forall_elim_list (map #2 params) o Thm.assume;
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
   121
    val subgoals = map lift (Drule.strip_imp_prems (Thm.cprop_of th));
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
   122
    val th' = fold (Thm.elim_implies o unlift) subgoals th;
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
   123
  in (subgoals, th') end;
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
   124
32213
f1b166317ae2 added focus_params/FOCUS_PARAMS, which focus on the parameter prefix only;
wenzelm
parents: 32210
diff changeset
   125
fun retrofit ctxt1 ctxt0 params asms i st1 st0 =
32200
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
   126
  let
32284
d8ee8a956f19 retrofit: more precise handling of locally introduced schematic variables (avoid zero_var_indexes);
wenzelm
parents: 32283
diff changeset
   127
    val idx = Thm.maxidx_of st0 + 1;
32200
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
   128
    val ps = map #2 params;
32284
d8ee8a956f19 retrofit: more precise handling of locally introduced schematic variables (avoid zero_var_indexes);
wenzelm
parents: 32283
diff changeset
   129
    val ((subgoal_inst, st2), ctxt2) = lift_import idx ps st1 ctxt1;
32213
f1b166317ae2 added focus_params/FOCUS_PARAMS, which focus on the parameter prefix only;
wenzelm
parents: 32210
diff changeset
   130
    val (subgoals, st3) = lift_subgoals params asms st2;
f1b166317ae2 added focus_params/FOCUS_PARAMS, which focus on the parameter prefix only;
wenzelm
parents: 32210
diff changeset
   131
    val result = st3
f1b166317ae2 added focus_params/FOCUS_PARAMS, which focus on the parameter prefix only;
wenzelm
parents: 32210
diff changeset
   132
      |> Goal.conclude
32200
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
   133
      |> Drule.implies_intr_list asms
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
   134
      |> Drule.forall_intr_list ps
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
   135
      |> Drule.implies_intr_list subgoals
32284
d8ee8a956f19 retrofit: more precise handling of locally introduced schematic variables (avoid zero_var_indexes);
wenzelm
parents: 32283
diff changeset
   136
      |> fold_rev (Thm.forall_intr o #1) subgoal_inst
d8ee8a956f19 retrofit: more precise handling of locally introduced schematic variables (avoid zero_var_indexes);
wenzelm
parents: 32283
diff changeset
   137
      |> fold (Thm.forall_elim o #2) subgoal_inst
d8ee8a956f19 retrofit: more precise handling of locally introduced schematic variables (avoid zero_var_indexes);
wenzelm
parents: 32283
diff changeset
   138
      |> Thm.adjust_maxidx_thm idx
d8ee8a956f19 retrofit: more precise handling of locally introduced schematic variables (avoid zero_var_indexes);
wenzelm
parents: 32283
diff changeset
   139
      |> singleton (Variable.export ctxt2 ctxt0);
52223
5bb6ae8acb87 tuned signature -- more explicit flags for low-level Thm.bicompose;
wenzelm
parents: 49845
diff changeset
   140
  in
58950
d07464875dd4 optional proof context for unify operations, for the sake of proper local options;
wenzelm
parents: 54883
diff changeset
   141
    Thm.bicompose (SOME ctxt0) {flatten = true, match = false, incremented = false}
52223
5bb6ae8acb87 tuned signature -- more explicit flags for low-level Thm.bicompose;
wenzelm
parents: 49845
diff changeset
   142
      (false, result, Thm.nprems_of st1) i st0
5bb6ae8acb87 tuned signature -- more explicit flags for low-level Thm.bicompose;
wenzelm
parents: 49845
diff changeset
   143
  end;
32200
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
   144
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
   145
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
   146
(* tacticals *)
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
   147
32281
750101435f60 focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
wenzelm
parents: 32213
diff changeset
   148
fun GEN_FOCUS flags tac ctxt i st =
20210
8fe4ae4360eb Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff changeset
   149
  if Thm.nprems_of st < i then Seq.empty
8fe4ae4360eb Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff changeset
   150
  else
32329
1527ff8c2dfb SUBPROOF: recovered Goal.check_finished;
wenzelm
parents: 32284
diff changeset
   151
    let val (args as {context = ctxt', params, asms, ...}, st') = gen_focus flags ctxt i st;
1527ff8c2dfb SUBPROOF: recovered Goal.check_finished;
wenzelm
parents: 32284
diff changeset
   152
    in Seq.lifts (retrofit ctxt' ctxt params asms i) (tac args st') st end;
32200
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
   153
32281
750101435f60 focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
wenzelm
parents: 32213
diff changeset
   154
val FOCUS_PARAMS = GEN_FOCUS (false, false);
60626
9eefb9f39021 clarified prems: full subgoal is imported in any case, to avoid remaining schematic variables;
wenzelm
parents: 60625
diff changeset
   155
val FOCUS_PARAMS_FIXED = GEN_FOCUS (false, true);
32281
750101435f60 focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
wenzelm
parents: 32213
diff changeset
   156
val FOCUS_PREMS = GEN_FOCUS (true, false);
750101435f60 focus: more precise treatment of schematic variables, only the "visible" part of the text is fixed;
wenzelm
parents: 32213
diff changeset
   157
val FOCUS = GEN_FOCUS (true, true);
32213
f1b166317ae2 added focus_params/FOCUS_PARAMS, which focus on the parameter prefix only;
wenzelm
parents: 32210
diff changeset
   158
49845
9b19c0e81166 tuned signature;
wenzelm
parents: 42495
diff changeset
   159
fun SUBPROOF tac ctxt = FOCUS (Seq.map (Goal.check_finished ctxt) oo tac) ctxt;
20210
8fe4ae4360eb Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff changeset
   160
60623
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   161
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   162
(* Isar subgoal command *)
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   163
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   164
local
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   165
60626
9eefb9f39021 clarified prems: full subgoal is imported in any case, to avoid remaining schematic variables;
wenzelm
parents: 60625
diff changeset
   166
fun gen_subgoal prep_atts raw_result_binding raw_prems_binding params state =
60623
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   167
  let
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   168
    val _ = Proof.assert_backward state;
60627
5d13babbb3f6 split multi-goals as usual (outermost Pure.conjunction only);
wenzelm
parents: 60626
diff changeset
   169
5d13babbb3f6 split multi-goals as usual (outermost Pure.conjunction only);
wenzelm
parents: 60626
diff changeset
   170
    val state1 = state |> Proof.refine_insert [];
5d13babbb3f6 split multi-goals as usual (outermost Pure.conjunction only);
wenzelm
parents: 60626
diff changeset
   171
    val {context = ctxt, facts = facts, goal = st} = Proof.raw_goal state1;
60623
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   172
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   173
    val result_binding = apsnd (map (prep_atts ctxt)) raw_result_binding;
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   174
    val (prems_binding, do_prems) =
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   175
      (case raw_prems_binding of
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   176
        SOME (b, raw_atts) => ((b, map (prep_atts ctxt) raw_atts), true)
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   177
      | NONE => ((Binding.empty, []), false));
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   178
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   179
    fun check_param (b, mx) = ignore (Proof_Context.cert_var (b, NONE, mx) ctxt);
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   180
    val _ = List.app check_param params;
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   181
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   182
    val _ = if Thm.no_prems st then error "No subgoals!" else ();
60626
9eefb9f39021 clarified prems: full subgoal is imported in any case, to avoid remaining schematic variables;
wenzelm
parents: 60625
diff changeset
   183
    val subgoal_focus =
9eefb9f39021 clarified prems: full subgoal is imported in any case, to avoid remaining schematic variables;
wenzelm
parents: 60625
diff changeset
   184
      #1 ((if do_prems then focus else focus_params_fixed) ctxt 1 st);
60623
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   185
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   186
    fun after_qed (ctxt'', [[result]]) =
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   187
      Proof.end_block #> (fn state' =>
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   188
        let
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   189
          val ctxt' = Proof.context_of state';
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   190
          val results' =
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   191
            Proof_Context.export ctxt'' ctxt' (Conjunction.elim_conjunctions result);
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   192
        in
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   193
          state'
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   194
          |> Proof.refine_primitive (fn _ => fn _ =>
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   195
            retrofit ctxt'' ctxt' (#params subgoal_focus) (#asms subgoal_focus) 1
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   196
              (Goal.protect 0 result) st
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   197
            |> Seq.hd)
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   198
          |> Proof.map_context
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   199
            (#2 o Proof_Context.note_thmss "" [(result_binding, [(results', [])])])
60625
f64658887a05 proper state after qed;
wenzelm
parents: 60623
diff changeset
   200
        end)
f64658887a05 proper state after qed;
wenzelm
parents: 60623
diff changeset
   201
      #> Proof.reset_facts
f64658887a05 proper state after qed;
wenzelm
parents: 60623
diff changeset
   202
      #> Proof.enter_backward;
60623
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   203
  in
60627
5d13babbb3f6 split multi-goals as usual (outermost Pure.conjunction only);
wenzelm
parents: 60626
diff changeset
   204
    state1
60623
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   205
    |> Proof.enter_forward
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   206
    |> Proof.using_facts []
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   207
    |> Proof.begin_block
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   208
    |> Proof.map_context (fn _ =>
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   209
      #context subgoal_focus
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   210
      |> Proof_Context.note_thmss "" [(prems_binding, [(#prems subgoal_focus, [])])] |> #2)
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   211
    |> Proof.internal_goal (K (K ())) (Proof_Context.get_mode ctxt) true "subgoal"
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   212
        NONE after_qed [] [] [(Thm.empty_binding, [(Thm.term_of (#concl subgoal_focus), [])])] |> #2
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   213
    |> Proof.using_facts facts
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   214
    |> pair subgoal_focus
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   215
  end;
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   216
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   217
in
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   218
60626
9eefb9f39021 clarified prems: full subgoal is imported in any case, to avoid remaining schematic variables;
wenzelm
parents: 60625
diff changeset
   219
val subgoal = gen_subgoal Attrib.attribute;
9eefb9f39021 clarified prems: full subgoal is imported in any case, to avoid remaining schematic variables;
wenzelm
parents: 60625
diff changeset
   220
val subgoal_cmd = gen_subgoal Attrib.attribute_cmd;
60623
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   221
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   222
end;
be39fe6c5620 support for subgoal focus command;
wenzelm
parents: 60367
diff changeset
   223
20210
8fe4ae4360eb Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff changeset
   224
end;
8fe4ae4360eb Tactical operations depending on local subgoal structure.
wenzelm
parents:
diff changeset
   225
32200
2bd8ab91a426 advanced retrofit, which allows new subgoals and variables;
wenzelm
parents: 31794
diff changeset
   226
val SUBPROOF = Subgoal.SUBPROOF;