src/Pure/Isar/obtain.ML
author haftmann
Sat, 19 Jul 2025 18:41:55 +0200
changeset 82886 8d1e295aab70
parent 80910 406a85a25189
permissions -rw-r--r--
clarified name and status of auxiliary operation
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
7674
99305245f6bd The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff changeset
     1
(*  Title:      Pure/Isar/obtain.ML
99305245f6bd The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff changeset
     2
    Author:     Markus Wenzel, TU Muenchen
99305245f6bd The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff changeset
     3
60452
3a0d57f1d6ef tuned comments;
wenzelm
parents: 60451
diff changeset
     4
Generalized existence and cases rules within Isar proof text.
8094
62b45a2e6ecb ObtainFun;
wenzelm
parents: 7923
diff changeset
     5
*)
7674
99305245f6bd The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff changeset
     6
99305245f6bd The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff changeset
     7
signature OBTAIN =
99305245f6bd The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff changeset
     8
sig
74365
b49bd5d9041f improper proof command 'guess' moved to separate theory "Pure-ex.Guess";
wenzelm
parents: 74282
diff changeset
     9
  val obtain_export: Proof.context -> thm -> cterm list -> Assumption.export
60477
051b200f7578 improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents: 60475
diff changeset
    10
  val obtain_thesis: Proof.context -> ((string * typ) * term) * Proof.context
60444
9945378d1ee7 tuned signature;
wenzelm
parents: 60415
diff changeset
    11
  val obtains_attributes: ('typ, 'term) Element.obtain list -> attribute list
63019
80ef19b51493 prefer internal attribute source;
wenzelm
parents: 61841
diff changeset
    12
  val obtains_attribs: ('typ, 'term) Element.obtain list -> Token.src list
60448
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
    13
  val read_obtains: Proof.context -> term -> Element.obtains -> (binding * term) list
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
    14
  val cert_obtains: Proof.context -> term -> Element.obtains_i -> (binding * term) list
60477
051b200f7578 improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents: 60475
diff changeset
    15
  val parse_obtains: Proof.context -> term -> Element.obtains -> (binding * term) list
60448
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
    16
  val consider: Element.obtains_i -> bool -> Proof.state -> Proof.state
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
    17
  val consider_cmd: Element.obtains -> bool -> Proof.state -> Proof.state
60453
ba9085f7433f clarified 'obtain', using structured 'have' statement;
wenzelm
parents: 60452
diff changeset
    18
  val obtain: binding -> (binding * typ option * mixfix) list ->
63059
3f577308551e 'obtain' supports structured statements (similar to 'define');
wenzelm
parents: 63057
diff changeset
    19
    (binding * typ option * mixfix) list -> (term * term list) list list ->
36323
655e2d74de3a modernized naming conventions of main Isar proof elements;
wenzelm
parents: 35845
diff changeset
    20
    (Thm.binding * (term * term list) list) list -> bool -> Proof.state -> Proof.state
60453
ba9085f7433f clarified 'obtain', using structured 'have' statement;
wenzelm
parents: 60452
diff changeset
    21
  val obtain_cmd: binding -> (binding * string option * mixfix) list ->
63059
3f577308551e 'obtain' supports structured statements (similar to 'define');
wenzelm
parents: 63057
diff changeset
    22
    (binding * string option * mixfix) list -> (string * string list) list list ->
30211
556d1810cdad Thm.binding;
wenzelm
parents: 29581
diff changeset
    23
    (Attrib.binding * (string * string list) list) list -> bool -> Proof.state -> Proof.state
74365
b49bd5d9041f improper proof command 'guess' moved to separate theory "Pure-ex.Guess";
wenzelm
parents: 74282
diff changeset
    24
  val check_result: Proof.context -> term -> thm -> thm
20308
ddb7e7129481 added tactical result;
wenzelm
parents: 20224
diff changeset
    25
  val result: (Proof.context -> tactic) -> thm list -> Proof.context ->
32199
82c4c570310a Variable.focus: named parameters;
wenzelm
parents: 32091
diff changeset
    26
    ((string * cterm) list * thm list) * Proof.context
7674
99305245f6bd The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff changeset
    27
end;
99305245f6bd The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff changeset
    28
10379
93630e0c5ae9 improved handling of "that": insert into goal, only declare as Pure "intro";
wenzelm
parents: 9481
diff changeset
    29
structure Obtain: OBTAIN =
7674
99305245f6bd The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff changeset
    30
struct
99305245f6bd The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff changeset
    31
60452
3a0d57f1d6ef tuned comments;
wenzelm
parents: 60451
diff changeset
    32
(** specification elements **)
3a0d57f1d6ef tuned comments;
wenzelm
parents: 60451
diff changeset
    33
3a0d57f1d6ef tuned comments;
wenzelm
parents: 60451
diff changeset
    34
(* obtain_export *)
18670
c3f445b92aff uniform handling of fixes;
wenzelm
parents: 18643
diff changeset
    35
18870
020e242c02a0 tuned comments;
wenzelm
parents: 18769
diff changeset
    36
(*
18897
b31293969d4f obtain(_i): optional name for 'that';
wenzelm
parents: 18870
diff changeset
    37
  [x, A x]
b31293969d4f obtain(_i): optional name for 'that';
wenzelm
parents: 18870
diff changeset
    38
     :
b31293969d4f obtain(_i): optional name for 'that';
wenzelm
parents: 18870
diff changeset
    39
     B
b31293969d4f obtain(_i): optional name for 'that';
wenzelm
parents: 18870
diff changeset
    40
  --------
b31293969d4f obtain(_i): optional name for 'that';
wenzelm
parents: 18870
diff changeset
    41
     B
18870
020e242c02a0 tuned comments;
wenzelm
parents: 18769
diff changeset
    42
*)
21686
4f5f6c685ab4 export: added explicit term operation;
wenzelm
parents: 21605
diff changeset
    43
fun eliminate_term ctxt xs tm =
4f5f6c685ab4 export: added explicit term operation;
wenzelm
parents: 21605
diff changeset
    44
  let
4f5f6c685ab4 export: added explicit term operation;
wenzelm
parents: 21605
diff changeset
    45
    val vs = map (dest_Free o Thm.term_of) xs;
4f5f6c685ab4 export: added explicit term operation;
wenzelm
parents: 21605
diff changeset
    46
    val bads = Term.fold_aterms (fn t as Free v =>
4f5f6c685ab4 export: added explicit term operation;
wenzelm
parents: 21605
diff changeset
    47
      if member (op =) vs v then insert (op aconv) t else I | _ => I) tm [];
4f5f6c685ab4 export: added explicit term operation;
wenzelm
parents: 21605
diff changeset
    48
    val _ = null bads orelse
4f5f6c685ab4 export: added explicit term operation;
wenzelm
parents: 21605
diff changeset
    49
      error ("Result contains obtained parameters: " ^
80910
406a85a25189 clarified signature: more explicit operations;
wenzelm
parents: 77908
diff changeset
    50
        implode_space (map (Syntax.string_of_term ctxt) bads));
21686
4f5f6c685ab4 export: added explicit term operation;
wenzelm
parents: 21605
diff changeset
    51
  in tm end;
4f5f6c685ab4 export: added explicit term operation;
wenzelm
parents: 21605
diff changeset
    52
60387
76359ff1090f more careful treatment of term bindings in 'obtain' proof body;
wenzelm
parents: 60383
diff changeset
    53
fun eliminate ctxt rule xs As thm =
9468
9adbcf6375c1 turned into plain context element;
wenzelm
parents: 9293
diff changeset
    54
  let
60387
76359ff1090f more careful treatment of term bindings in 'obtain' proof body;
wenzelm
parents: 60383
diff changeset
    55
    val _ = eliminate_term ctxt xs (Thm.full_prop_of thm);
76359ff1090f more careful treatment of term bindings in 'obtain' proof body;
wenzelm
parents: 60383
diff changeset
    56
    val _ = Object_Logic.is_judgment ctxt (Thm.concl_of thm) orelse
20308
ddb7e7129481 added tactical result;
wenzelm
parents: 20224
diff changeset
    57
      error "Conclusion in obtained context must be object-logic judgment";
ddb7e7129481 added tactical result;
wenzelm
parents: 20224
diff changeset
    58
60387
76359ff1090f more careful treatment of term bindings in 'obtain' proof body;
wenzelm
parents: 60383
diff changeset
    59
    val ((_, [thm']), ctxt') = Variable.import true [thm] ctxt;
60949
ccbf9379e355 added Thm.chyps_of;
wenzelm
parents: 60695
diff changeset
    60
    val prems = Drule.strip_imp_prems (Thm.cprop_of thm');
9468
9adbcf6375c1 turned into plain context element;
wenzelm
parents: 9293
diff changeset
    61
  in
20308
ddb7e7129481 added tactical result;
wenzelm
parents: 20224
diff changeset
    62
    ((Drule.implies_elim_list thm' (map Thm.assume prems)
60315
c08adefc98ea more explicit context;
wenzelm
parents: 59970
diff changeset
    63
        |> Drule.implies_intr_list (map (Drule.norm_hhf_cterm ctxt') As)
20308
ddb7e7129481 added tactical result;
wenzelm
parents: 20224
diff changeset
    64
        |> Drule.forall_intr_list xs)
ddb7e7129481 added tactical result;
wenzelm
parents: 20224
diff changeset
    65
      COMP rule)
ddb7e7129481 added tactical result;
wenzelm
parents: 20224
diff changeset
    66
    |> Drule.implies_intr_list prems
60387
76359ff1090f more careful treatment of term bindings in 'obtain' proof body;
wenzelm
parents: 60383
diff changeset
    67
    |> singleton (Variable.export ctxt' ctxt)
9468
9adbcf6375c1 turned into plain context element;
wenzelm
parents: 9293
diff changeset
    68
  end;
9adbcf6375c1 turned into plain context element;
wenzelm
parents: 9293
diff changeset
    69
21686
4f5f6c685ab4 export: added explicit term operation;
wenzelm
parents: 21605
diff changeset
    70
fun obtain_export ctxt rule xs _ As =
4f5f6c685ab4 export: added explicit term operation;
wenzelm
parents: 21605
diff changeset
    71
  (eliminate ctxt rule xs As, eliminate_term ctxt xs);
4f5f6c685ab4 export: added explicit term operation;
wenzelm
parents: 21605
diff changeset
    72
9468
9adbcf6375c1 turned into plain context element;
wenzelm
parents: 9293
diff changeset
    73
60448
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
    74
(* result declaration *)
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
    75
63019
80ef19b51493 prefer internal attribute source;
wenzelm
parents: 61841
diff changeset
    76
fun case_names (obtains: ('typ, 'term) Element.obtain list) =
80ef19b51493 prefer internal attribute source;
wenzelm
parents: 61841
diff changeset
    77
  obtains |> map_index (fn (i, (b, _)) =>
80ef19b51493 prefer internal attribute source;
wenzelm
parents: 61841
diff changeset
    78
    if Binding.is_empty b then string_of_int (i + 1) else Name_Space.base_name b);
80ef19b51493 prefer internal attribute source;
wenzelm
parents: 61841
diff changeset
    79
80ef19b51493 prefer internal attribute source;
wenzelm
parents: 61841
diff changeset
    80
fun obtains_attributes obtains =
80ef19b51493 prefer internal attribute source;
wenzelm
parents: 61841
diff changeset
    81
  [Rule_Cases.consumes (~ (length obtains)), Rule_Cases.case_names (case_names obtains)];
80ef19b51493 prefer internal attribute source;
wenzelm
parents: 61841
diff changeset
    82
80ef19b51493 prefer internal attribute source;
wenzelm
parents: 61841
diff changeset
    83
fun obtains_attribs obtains =
80ef19b51493 prefer internal attribute source;
wenzelm
parents: 61841
diff changeset
    84
  [Attrib.consumes (~ (length obtains)), Attrib.case_names (case_names obtains)];
60444
9945378d1ee7 tuned signature;
wenzelm
parents: 60415
diff changeset
    85
9945378d1ee7 tuned signature;
wenzelm
parents: 60415
diff changeset
    86
60448
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
    87
(* obtain thesis *)
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
    88
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
    89
fun obtain_thesis ctxt =
60446
64f48e7f921f support to parse obtain clause without type-checking yet;
wenzelm
parents: 60444
diff changeset
    90
  let
64f48e7f921f support to parse obtain clause without type-checking yet;
wenzelm
parents: 60444
diff changeset
    91
    val ([x], ctxt') =
64f48e7f921f support to parse obtain clause without type-checking yet;
wenzelm
parents: 60444
diff changeset
    92
      Proof_Context.add_fixes [(Binding.name Auto_Bind.thesisN, NONE, NoSyn)] ctxt;
64f48e7f921f support to parse obtain clause without type-checking yet;
wenzelm
parents: 60444
diff changeset
    93
    val t = Object_Logic.fixed_judgment ctxt x;
64f48e7f921f support to parse obtain clause without type-checking yet;
wenzelm
parents: 60444
diff changeset
    94
    val v = dest_Free (Object_Logic.drop_judgment ctxt t);
64f48e7f921f support to parse obtain clause without type-checking yet;
wenzelm
parents: 60444
diff changeset
    95
  in ((v, t), ctxt') end;
64f48e7f921f support to parse obtain clause without type-checking yet;
wenzelm
parents: 60444
diff changeset
    96
60448
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
    97
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
    98
(* obtain clauses *)
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
    99
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
   100
local
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
   101
60477
051b200f7578 improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents: 60475
diff changeset
   102
val mk_all_external = Logic.all_constraint o Variable.default_type;
60481
09b04c815fdb more robust: variables need not occur in body;
wenzelm
parents: 60477
diff changeset
   103
09b04c815fdb more robust: variables need not occur in body;
wenzelm
parents: 60477
diff changeset
   104
fun mk_all_internal ctxt (y, z) t =
09b04c815fdb more robust: variables need not occur in body;
wenzelm
parents: 60477
diff changeset
   105
  let
74266
612b7e0d6721 clarified signature;
wenzelm
parents: 74232
diff changeset
   106
    val frees = Frees.build (t |> Term.fold_aterms (fn Free v => Frees.add_set v | _ => I));
60481
09b04c815fdb more robust: variables need not occur in body;
wenzelm
parents: 60477
diff changeset
   107
    val T =
77908
a6bd716a6124 tuned: concise combinators instead of bulky case-expressions;
wenzelm
parents: 74365
diff changeset
   108
      Frees.get_first (fn ((x, T), _) => if x = z then SOME T else NONE) frees
a6bd716a6124 tuned: concise combinators instead of bulky case-expressions;
wenzelm
parents: 74365
diff changeset
   109
      |> \<^if_none>\<open>the_default dummyT (Variable.default_type ctxt z)\<close>;
60477
051b200f7578 improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents: 60475
diff changeset
   110
  in Logic.all_const T $ Term.lambda_name (y, Free (z, T)) t end;
051b200f7578 improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents: 60475
diff changeset
   111
051b200f7578 improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents: 60475
diff changeset
   112
fun prepare_clause prep_var parse_prop mk_all ctxt thesis raw_vars raw_props =
60446
64f48e7f921f support to parse obtain clause without type-checking yet;
wenzelm
parents: 60444
diff changeset
   113
  let
60477
051b200f7578 improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents: 60475
diff changeset
   114
    val ((xs', vars), ctxt') = ctxt
051b200f7578 improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents: 60475
diff changeset
   115
      |> fold_map prep_var raw_vars
051b200f7578 improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents: 60475
diff changeset
   116
      |-> (fn vars => Proof_Context.add_fixes vars ##>> pair vars);
60446
64f48e7f921f support to parse obtain clause without type-checking yet;
wenzelm
parents: 60444
diff changeset
   117
    val xs = map (Variable.check_name o #1) vars;
64f48e7f921f support to parse obtain clause without type-checking yet;
wenzelm
parents: 60444
diff changeset
   118
  in
60448
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
   119
    Logic.list_implies (map (parse_prop ctxt') raw_props, thesis)
60477
051b200f7578 improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents: 60475
diff changeset
   120
    |> fold_rev (mk_all ctxt') (xs ~~ xs')
60446
64f48e7f921f support to parse obtain clause without type-checking yet;
wenzelm
parents: 60444
diff changeset
   121
  end;
64f48e7f921f support to parse obtain clause without type-checking yet;
wenzelm
parents: 60444
diff changeset
   122
60477
051b200f7578 improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents: 60475
diff changeset
   123
fun prepare_obtains prep_clause check_terms
051b200f7578 improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents: 60475
diff changeset
   124
    ctxt thesis (raw_obtains: ('typ, 'term) Element.obtain list) =
60448
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
   125
  let
60477
051b200f7578 improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents: 60475
diff changeset
   126
    val clauses = raw_obtains
051b200f7578 improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents: 60475
diff changeset
   127
      |> map (fn (_, (raw_vars, raw_props)) => prep_clause ctxt thesis raw_vars raw_props)
051b200f7578 improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents: 60475
diff changeset
   128
      |> check_terms ctxt;
051b200f7578 improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents: 60475
diff changeset
   129
  in map fst raw_obtains ~~ clauses end;
60448
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
   130
60477
051b200f7578 improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents: 60475
diff changeset
   131
val parse_clause = prepare_clause Proof_Context.read_var Syntax.parse_prop mk_all_external;
051b200f7578 improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents: 60475
diff changeset
   132
val cert_clause = prepare_clause Proof_Context.cert_var (K I) mk_all_internal;
60448
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
   133
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
   134
in
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
   135
60477
051b200f7578 improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents: 60475
diff changeset
   136
val read_obtains = prepare_obtains parse_clause Syntax.check_terms;
051b200f7578 improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents: 60475
diff changeset
   137
val cert_obtains = prepare_obtains cert_clause (K I);
051b200f7578 improved treatment of Element.Obtains via Expression.prepare_stmt;
wenzelm
parents: 60475
diff changeset
   138
val parse_obtains = prepare_obtains parse_clause (K I);
60448
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
   139
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
   140
end;
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
   141
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
   142
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
   143
60451
1f2b29f78439 clarified 'consider', using structured 'have' statement;
wenzelm
parents: 60448
diff changeset
   144
(** consider: generalized elimination and cases rule **)
1f2b29f78439 clarified 'consider', using structured 'have' statement;
wenzelm
parents: 60448
diff changeset
   145
1f2b29f78439 clarified 'consider', using structured 'have' statement;
wenzelm
parents: 60448
diff changeset
   146
(*
67721
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 63352
diff changeset
   147
  consider (a) x where "A x" | (b) y where "B y" | ... \<equiv>
60451
1f2b29f78439 clarified 'consider', using structured 'have' statement;
wenzelm
parents: 60448
diff changeset
   148
1f2b29f78439 clarified 'consider', using structured 'have' statement;
wenzelm
parents: 60448
diff changeset
   149
  have thesis
67721
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 63352
diff changeset
   150
    if a [intro?]: "\<And>x. A x \<Longrightarrow> thesis"
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 63352
diff changeset
   151
    and b [intro?]: "\<And>y. B y \<Longrightarrow> thesis"
60451
1f2b29f78439 clarified 'consider', using structured 'have' statement;
wenzelm
parents: 60448
diff changeset
   152
    and ...
1f2b29f78439 clarified 'consider', using structured 'have' statement;
wenzelm
parents: 60448
diff changeset
   153
    for thesis
1f2b29f78439 clarified 'consider', using structured 'have' statement;
wenzelm
parents: 60448
diff changeset
   154
    apply (insert that)
1f2b29f78439 clarified 'consider', using structured 'have' statement;
wenzelm
parents: 60448
diff changeset
   155
*)
60448
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
   156
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
   157
local
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
   158
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
   159
fun gen_consider prep_obtains raw_obtains int state =
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
   160
  let
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
   161
    val _ = Proof.assert_forward_or_chain state;
60453
ba9085f7433f clarified 'obtain', using structured 'have' statement;
wenzelm
parents: 60452
diff changeset
   162
    val ctxt = Proof.context_of state;
60451
1f2b29f78439 clarified 'consider', using structured 'have' statement;
wenzelm
parents: 60448
diff changeset
   163
1f2b29f78439 clarified 'consider', using structured 'have' statement;
wenzelm
parents: 60448
diff changeset
   164
    val ((_, thesis), thesis_ctxt) = obtain_thesis ctxt;
60448
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
   165
    val obtains = prep_obtains thesis_ctxt thesis raw_obtains;
60456
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60454
diff changeset
   166
    val atts = Rule_Cases.cases_open :: obtains_attributes raw_obtains;
60448
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
   167
  in
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
   168
    state
60555
51a6997b1384 support 'when' statement, which corresponds to 'presume';
wenzelm
parents: 60481
diff changeset
   169
    |> Proof.have true NONE (K I)
60451
1f2b29f78439 clarified 'consider', using structured 'have' statement;
wenzelm
parents: 60448
diff changeset
   170
      [(Binding.name Auto_Bind.thesisN, NONE, NoSyn)]
1f2b29f78439 clarified 'consider', using structured 'have' statement;
wenzelm
parents: 60448
diff changeset
   171
      (map (fn (a, A) => ((a, [Context_Rules.intro_query NONE]), [(A, [])])) obtains)
60456
323b15b5af73 open parameters for 'consider' rule;
wenzelm
parents: 60454
diff changeset
   172
      [((Binding.empty, atts), [(thesis, [])])] int
60461
22995ec9fefd tuned signature;
wenzelm
parents: 60456
diff changeset
   173
    |-> Proof.refine_insert
60448
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
   174
  end;
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
   175
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
   176
in
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
   177
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
   178
val consider = gen_consider cert_obtains;
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
   179
val consider_cmd = gen_consider read_obtains;
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
   180
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
   181
end;
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
   182
60446
64f48e7f921f support to parse obtain clause without type-checking yet;
wenzelm
parents: 60444
diff changeset
   183
60444
9945378d1ee7 tuned signature;
wenzelm
parents: 60415
diff changeset
   184
60452
3a0d57f1d6ef tuned comments;
wenzelm
parents: 60451
diff changeset
   185
(** obtain: augmented context based on generalized existence rule **)
3a0d57f1d6ef tuned comments;
wenzelm
parents: 60451
diff changeset
   186
3a0d57f1d6ef tuned comments;
wenzelm
parents: 60451
diff changeset
   187
(*
67721
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 63352
diff changeset
   188
  obtain (a) x where "A x" <proof> \<equiv>
60452
3a0d57f1d6ef tuned comments;
wenzelm
parents: 60451
diff changeset
   189
67721
5348bea4accd eliminated ASCII syntax from Pure bootstrap;
wenzelm
parents: 63352
diff changeset
   190
  have thesis if a [intro?]: "\<And>x. A x \<Longrightarrow> thesis" for thesis
60453
ba9085f7433f clarified 'obtain', using structured 'have' statement;
wenzelm
parents: 60452
diff changeset
   191
    apply (insert that)
ba9085f7433f clarified 'obtain', using structured 'have' statement;
wenzelm
parents: 60452
diff changeset
   192
    <proof>
60452
3a0d57f1d6ef tuned comments;
wenzelm
parents: 60451
diff changeset
   193
  fix x assm <<obtain_export>> "A x"
3a0d57f1d6ef tuned comments;
wenzelm
parents: 60451
diff changeset
   194
*)
17858
bc4db8cfd92f added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents: 17357
diff changeset
   195
bc4db8cfd92f added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents: 17357
diff changeset
   196
local
8094
62b45a2e6ecb ObtainFun;
wenzelm
parents: 7923
diff changeset
   197
63059
3f577308551e 'obtain' supports structured statements (similar to 'define');
wenzelm
parents: 63057
diff changeset
   198
fun gen_obtain prep_stmt prep_att that_binding raw_decls raw_fixes raw_prems raw_concls int state =
7674
99305245f6bd The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff changeset
   199
  let
9468
9adbcf6375c1 turned into plain context element;
wenzelm
parents: 9293
diff changeset
   200
    val _ = Proof.assert_forward_or_chain state;
7674
99305245f6bd The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff changeset
   201
63037
b8b672f70d76 clarified modules;
wenzelm
parents: 63019
diff changeset
   202
    val ((_, thesis), thesis_ctxt) = obtain_thesis (Proof.context_of state);
63059
3f577308551e 'obtain' supports structured statements (similar to 'define');
wenzelm
parents: 63057
diff changeset
   203
70734
31364e70ff3e clarified signature;
wenzelm
parents: 70729
diff changeset
   204
    val ({vars, propss, binds, result_binds, ...}, params_ctxt) =
63059
3f577308551e 'obtain' supports structured statements (similar to 'define');
wenzelm
parents: 63057
diff changeset
   205
      prep_stmt (raw_decls @ raw_fixes) (raw_prems @ map #2 raw_concls) thesis_ctxt;
3f577308551e 'obtain' supports structured statements (similar to 'define');
wenzelm
parents: 63057
diff changeset
   206
    val (decls, fixes) = chop (length raw_decls) vars ||> map #2;
3f577308551e 'obtain' supports structured statements (similar to 'define');
wenzelm
parents: 63057
diff changeset
   207
    val (premss, conclss) = chop (length raw_prems) propss;
3f577308551e 'obtain' supports structured statements (similar to 'define');
wenzelm
parents: 63057
diff changeset
   208
    val propss' = (map o map) (Logic.close_prop fixes (flat premss)) conclss;
3f577308551e 'obtain' supports structured statements (similar to 'define');
wenzelm
parents: 63057
diff changeset
   209
63056
9b95ae9ec671 defs are closed, which leads to proper auto_bind_facts;
wenzelm
parents: 63037
diff changeset
   210
    val that_prop =
63059
3f577308551e 'obtain' supports structured statements (similar to 'define');
wenzelm
parents: 63057
diff changeset
   211
      Logic.list_rename_params (map (#1 o #2) decls)
3f577308551e 'obtain' supports structured statements (similar to 'define');
wenzelm
parents: 63057
diff changeset
   212
        (fold_rev (Logic.all o #2 o #2) decls (Logic.list_implies (flat propss', thesis)));
7674
99305245f6bd The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff changeset
   213
63059
3f577308551e 'obtain' supports structured statements (similar to 'define');
wenzelm
parents: 63057
diff changeset
   214
    val cparams = map (Thm.cterm_of params_ctxt o #2 o #2) decls;
60387
76359ff1090f more careful treatment of term bindings in 'obtain' proof body;
wenzelm
parents: 60383
diff changeset
   215
    val asms =
63059
3f577308551e 'obtain' supports structured statements (similar to 'define');
wenzelm
parents: 63057
diff changeset
   216
      map (fn ((b, raw_atts), _) => (b, map (prep_att params_ctxt) raw_atts)) raw_concls ~~
3f577308551e 'obtain' supports structured statements (similar to 'define');
wenzelm
parents: 63057
diff changeset
   217
      map (map (rpair [])) propss';
10464
b7b916a82dca tuned statement args;
wenzelm
parents: 10379
diff changeset
   218
60453
ba9085f7433f clarified 'obtain', using structured 'have' statement;
wenzelm
parents: 60452
diff changeset
   219
    fun after_qed (result_ctxt, results) state' =
ba9085f7433f clarified 'obtain', using structured 'have' statement;
wenzelm
parents: 60452
diff changeset
   220
      let val [rule] = Proof_Context.export result_ctxt (Proof.context_of state') (flat results) in
ba9085f7433f clarified 'obtain', using structured 'have' statement;
wenzelm
parents: 60452
diff changeset
   221
        state'
63059
3f577308551e 'obtain' supports structured statements (similar to 'define');
wenzelm
parents: 63057
diff changeset
   222
        |> Proof.fix (map #1 decls)
3f577308551e 'obtain' supports structured statements (similar to 'define');
wenzelm
parents: 63057
diff changeset
   223
        |> Proof.map_context (fold (Variable.bind_term o apsnd (Logic.close_term fixes)) binds)
61654
4a28eec739e9 support for structure statements in 'assume', 'presume';
wenzelm
parents: 61268
diff changeset
   224
        |> Proof.assm (obtain_export params_ctxt rule cparams) [] [] asms
60453
ba9085f7433f clarified 'obtain', using structured 'have' statement;
wenzelm
parents: 60452
diff changeset
   225
      end;
7674
99305245f6bd The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff changeset
   226
  in
8094
62b45a2e6ecb ObtainFun;
wenzelm
parents: 7923
diff changeset
   227
    state
60555
51a6997b1384 support 'when' statement, which corresponds to 'presume';
wenzelm
parents: 60481
diff changeset
   228
    |> Proof.have true NONE after_qed
60453
ba9085f7433f clarified 'obtain', using structured 'have' statement;
wenzelm
parents: 60452
diff changeset
   229
      [(Binding.name Auto_Bind.thesisN, NONE, NoSyn)]
ba9085f7433f clarified 'obtain', using structured 'have' statement;
wenzelm
parents: 60452
diff changeset
   230
      [((that_binding, [Context_Rules.intro_query NONE]), [(that_prop, [])])]
63352
4eaf35781b23 tuned signature;
wenzelm
parents: 63059
diff changeset
   231
      [(Binding.empty_atts, [(thesis, [])])] int
60461
22995ec9fefd tuned signature;
wenzelm
parents: 60456
diff changeset
   232
    |-> Proof.refine_insert
70734
31364e70ff3e clarified signature;
wenzelm
parents: 70729
diff changeset
   233
    |> Proof.map_context (fold Variable.bind_term result_binds)
7674
99305245f6bd The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff changeset
   234
  end;
99305245f6bd The 'obtain' language element -- achieves (eliminated) existential
wenzelm
parents:
diff changeset
   235
17858
bc4db8cfd92f added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents: 17357
diff changeset
   236
in
bc4db8cfd92f added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents: 17357
diff changeset
   237
63057
50802acac277 more uniform operations for structured statements;
wenzelm
parents: 63056
diff changeset
   238
val obtain = gen_obtain Proof_Context.cert_stmt (K I);
50802acac277 more uniform operations for structured statements;
wenzelm
parents: 63056
diff changeset
   239
val obtain_cmd = gen_obtain Proof_Context.read_stmt Attrib.attribute_cmd;
8094
62b45a2e6ecb ObtainFun;
wenzelm
parents: 7923
diff changeset
   240
62b45a2e6ecb ObtainFun;
wenzelm
parents: 7923
diff changeset
   241
end;
17858
bc4db8cfd92f added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents: 17357
diff changeset
   242
bc4db8cfd92f added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents: 17357
diff changeset
   243
bc4db8cfd92f added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents: 17357
diff changeset
   244
20308
ddb7e7129481 added tactical result;
wenzelm
parents: 20224
diff changeset
   245
(** tactical result **)
ddb7e7129481 added tactical result;
wenzelm
parents: 20224
diff changeset
   246
ddb7e7129481 added tactical result;
wenzelm
parents: 20224
diff changeset
   247
fun check_result ctxt thesis th =
ddb7e7129481 added tactical result;
wenzelm
parents: 20224
diff changeset
   248
  (case Thm.prems_of th of
ddb7e7129481 added tactical result;
wenzelm
parents: 20224
diff changeset
   249
    [prem] =>
ddb7e7129481 added tactical result;
wenzelm
parents: 20224
diff changeset
   250
      if Thm.concl_of th aconv thesis andalso
ddb7e7129481 added tactical result;
wenzelm
parents: 20224
diff changeset
   251
        Logic.strip_assums_concl prem aconv thesis then th
61268
abe08fb15a12 moved remaining display.ML to more_thm.ML;
wenzelm
parents: 60949
diff changeset
   252
      else error ("Guessed a different clause:\n" ^ Thm.string_of_thm ctxt th)
38875
c7a66b584147 tuned messages: discontinued spurious full-stops (messages are occasionally composed unexpectedly);
wenzelm
parents: 36323
diff changeset
   253
  | [] => error "Goal solved -- nothing guessed"
61268
abe08fb15a12 moved remaining display.ML to more_thm.ML;
wenzelm
parents: 60949
diff changeset
   254
  | _ => error ("Guess split into several cases:\n" ^ Thm.string_of_thm ctxt th));
20308
ddb7e7129481 added tactical result;
wenzelm
parents: 20224
diff changeset
   255
ddb7e7129481 added tactical result;
wenzelm
parents: 20224
diff changeset
   256
fun result tac facts ctxt =
ddb7e7129481 added tactical result;
wenzelm
parents: 20224
diff changeset
   257
  let
60448
78f3c67bc35e support for 'consider' command;
wenzelm
parents: 60446
diff changeset
   258
    val ((thesis_var, thesis), thesis_ctxt) = obtain_thesis ctxt;
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59616
diff changeset
   259
    val st = Goal.init (Thm.cterm_of ctxt thesis);
20308
ddb7e7129481 added tactical result;
wenzelm
parents: 20224
diff changeset
   260
    val rule =
61841
4d3527b94f2a more general types Proof.method / context_tactic;
wenzelm
parents: 61654
diff changeset
   261
      (case SINGLE (Method.insert_tac thesis_ctxt facts 1 THEN tac thesis_ctxt) st of
20308
ddb7e7129481 added tactical result;
wenzelm
parents: 20224
diff changeset
   262
        NONE => raise THM ("Obtain.result: tactic failed", 0, facts)
54883
dd04a8b654fc proper context for norm_hhf and derived operations;
wenzelm
parents: 52456
diff changeset
   263
      | SOME th =>
dd04a8b654fc proper context for norm_hhf and derived operations;
wenzelm
parents: 52456
diff changeset
   264
          check_result thesis_ctxt thesis (Raw_Simplifier.norm_hhf thesis_ctxt (Goal.conclude th)));
20308
ddb7e7129481 added tactical result;
wenzelm
parents: 20224
diff changeset
   265
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59616
diff changeset
   266
    val closed_rule = Thm.forall_intr (Thm.cterm_of ctxt (Free thesis_var)) rule;
31794
71af1fd6a5e4 renamed Variable.import_thms to Variable.import (back again cf. ed7aa5a350ef -- Alice is no longer supported);
wenzelm
parents: 30763
diff changeset
   267
    val ((_, [rule']), ctxt') = Variable.import false [closed_rule] ctxt;
59616
eb59c6968219 tuned -- more explicit use of context;
wenzelm
parents: 59573
diff changeset
   268
    val obtain_rule =
59621
291934bac95e Thm.cterm_of and Thm.ctyp_of operate on local context;
wenzelm
parents: 59616
diff changeset
   269
      Thm.forall_elim (Thm.cterm_of ctxt (Logic.varify_global (Free thesis_var))) rule';
60695
757549b4bbe6 Variable.focus etc.: optional bindings provided by user;
wenzelm
parents: 60642
diff changeset
   270
    val ((params, stmt), fix_ctxt) = Variable.focus_cterm NONE (Thm.cprem_of obtain_rule 1) ctxt';
20308
ddb7e7129481 added tactical result;
wenzelm
parents: 20224
diff changeset
   271
    val (prems, ctxt'') =
32199
82c4c570310a Variable.focus: named parameters;
wenzelm
parents: 32091
diff changeset
   272
      Assumption.add_assms (obtain_export fix_ctxt obtain_rule (map #2 params))
20308
ddb7e7129481 added tactical result;
wenzelm
parents: 20224
diff changeset
   273
        (Drule.strip_imp_prems stmt) fix_ctxt;
ddb7e7129481 added tactical result;
wenzelm
parents: 20224
diff changeset
   274
  in ((params, prems), ctxt'') end;
ddb7e7129481 added tactical result;
wenzelm
parents: 20224
diff changeset
   275
17858
bc4db8cfd92f added 'guess', which derives the obtained context from the course of reasoning;
wenzelm
parents: 17357
diff changeset
   276
end;