1(*  Title:      Pure/Isar/obtain.ML
2    Author:     Markus Wenzel, TU Muenchen
3
4Generalized existence and cases rules within Isar proof text.
5*)
6
7signature OBTAIN =
8sig
9  val obtain_thesis: Proof.context -> ((string * typ) * term) * Proof.context
10  val obtains_attributes: ('typ, 'term) Element.obtain list -> attribute list
11  val obtains_attribs: ('typ, 'term) Element.obtain list -> Token.src list
12  val read_obtains: Proof.context -> term -> Element.obtains -> (binding * term) list
13  val cert_obtains: Proof.context -> term -> Element.obtains_i -> (binding * term) list
14  val parse_obtains: Proof.context -> term -> Element.obtains -> (binding * term) list
15  val consider: Element.obtains_i -> bool -> Proof.state -> Proof.state
16  val consider_cmd: Element.obtains -> bool -> Proof.state -> Proof.state
17  val obtain: binding -> (binding * typ option * mixfix) list ->
18    (binding * typ option * mixfix) list -> (term * term list) list list ->
19    (Thm.binding * (term * term list) list) list -> bool -> Proof.state -> Proof.state
20  val obtain_cmd: binding -> (binding * string option * mixfix) list ->
21    (binding * string option * mixfix) list -> (string * string list) list list ->
22    (Attrib.binding * (string * string list) list) list -> bool -> Proof.state -> Proof.state
23  val result: (Proof.context -> tactic) -> thm list -> Proof.context ->
24    ((string * cterm) list * thm list) * Proof.context
25  val guess: (binding * typ option * mixfix) list -> bool -> Proof.state -> Proof.state
26  val guess_cmd: (binding * string option * mixfix) list -> bool -> Proof.state -> Proof.state
27end;
28
29structure Obtain: OBTAIN =
30struct
31
32(** specification elements **)
33
34(* obtain_export *)
35
36(*
37  [x, A x]
38     :
39     B
40  --------
41     B
42*)
43fun eliminate_term ctxt xs tm =
44  let
45    val vs = map (dest_Free o Thm.term_of) xs;
46    val bads = Term.fold_aterms (fn t as Free v =>
47      if member (op =) vs v then insert (op aconv) t else I | _ => I) tm [];
48    val _ = null bads orelse
49      error ("Result contains obtained parameters: " ^
50        space_implode " " (map (Syntax.string_of_term ctxt) bads));
51  in tm end;
52
53fun eliminate ctxt rule xs As thm =
54  let
55    val _ = eliminate_term ctxt xs (Thm.full_prop_of thm);
56    val _ = Object_Logic.is_judgment ctxt (Thm.concl_of thm) orelse
57      error "Conclusion in obtained context must be object-logic judgment";
58
59    val ((_, [thm']), ctxt') = Variable.import true [thm] ctxt;
60    val prems = Drule.strip_imp_prems (Thm.cprop_of thm');
61  in
62    ((Drule.implies_elim_list thm' (map Thm.assume prems)
63        |> Drule.implies_intr_list (map (Drule.norm_hhf_cterm ctxt') As)
64        |> Drule.forall_intr_list xs)
65      COMP rule)
66    |> Drule.implies_intr_list prems
67    |> singleton (Variable.export ctxt' ctxt)
68  end;
69
70fun obtain_export ctxt rule xs _ As =
71  (eliminate ctxt rule xs As, eliminate_term ctxt xs);
72
73
74(* result declaration *)
75
76fun case_names (obtains: ('typ, 'term) Element.obtain list) =
77  obtains |> map_index (fn (i, (b, _)) =>
78    if Binding.is_empty b then string_of_int (i + 1) else Name_Space.base_name b);
79
80fun obtains_attributes obtains =
81  [Rule_Cases.consumes (~ (length obtains)), Rule_Cases.case_names (case_names obtains)];
82
83fun obtains_attribs obtains =
84  [Attrib.consumes (~ (length obtains)), Attrib.case_names (case_names obtains)];
85
86
87(* obtain thesis *)
88
89fun obtain_thesis ctxt =
90  let
91    val ([x], ctxt') =
92      Proof_Context.add_fixes [(Binding.name Auto_Bind.thesisN, NONE, NoSyn)] ctxt;
93    val t = Object_Logic.fixed_judgment ctxt x;
94    val v = dest_Free (Object_Logic.drop_judgment ctxt t);
95  in ((v, t), ctxt') end;
96
97
98(* obtain clauses *)
99
100local
101
102val mk_all_external = Logic.all_constraint o Variable.default_type;
103
104fun mk_all_internal ctxt (y, z) t =
105  let
106    val T =
107      (case AList.lookup (op =) (Term.add_frees t []) z of
108        SOME T => T
109      | NONE => the_default dummyT (Variable.default_type ctxt z));
110  in Logic.all_const T $ Term.lambda_name (y, Free (z, T)) t end;
111
112fun prepare_clause prep_var parse_prop mk_all ctxt thesis raw_vars raw_props =
113  let
114    val ((xs', vars), ctxt') = ctxt
115      |> fold_map prep_var raw_vars
116      |-> (fn vars => Proof_Context.add_fixes vars ##>> pair vars);
117    val xs = map (Variable.check_name o #1) vars;
118  in
119    Logic.list_implies (map (parse_prop ctxt') raw_props, thesis)
120    |> fold_rev (mk_all ctxt') (xs ~~ xs')
121  end;
122
123fun prepare_obtains prep_clause check_terms
124    ctxt thesis (raw_obtains: ('typ, 'term) Element.obtain list) =
125  let
126    val clauses = raw_obtains
127      |> map (fn (_, (raw_vars, raw_props)) => prep_clause ctxt thesis raw_vars raw_props)
128      |> check_terms ctxt;
129  in map fst raw_obtains ~~ clauses end;
130
131val parse_clause = prepare_clause Proof_Context.read_var Syntax.parse_prop mk_all_external;
132val cert_clause = prepare_clause Proof_Context.cert_var (K I) mk_all_internal;
133
134in
135
136val read_obtains = prepare_obtains parse_clause Syntax.check_terms;
137val cert_obtains = prepare_obtains cert_clause (K I);
138val parse_obtains = prepare_obtains parse_clause (K I);
139
140end;
141
142
143
144(** consider: generalized elimination and cases rule **)
145
146(*
147  consider (a) x where "A x" | (b) y where "B y" | ... \<equiv>
148
149  have thesis
150    if a [intro?]: "\<And>x. A x \<Longrightarrow> thesis"
151    and b [intro?]: "\<And>y. B y \<Longrightarrow> thesis"
152    and ...
153    for thesis
154    apply (insert that)
155*)
156
157local
158
159fun gen_consider prep_obtains raw_obtains int state =
160  let
161    val _ = Proof.assert_forward_or_chain state;
162    val ctxt = Proof.context_of state;
163
164    val ((_, thesis), thesis_ctxt) = obtain_thesis ctxt;
165    val obtains = prep_obtains thesis_ctxt thesis raw_obtains;
166    val atts = Rule_Cases.cases_open :: obtains_attributes raw_obtains;
167  in
168    state
169    |> Proof.have true NONE (K I)
170      [(Binding.name Auto_Bind.thesisN, NONE, NoSyn)]
171      (map (fn (a, A) => ((a, [Context_Rules.intro_query NONE]), [(A, [])])) obtains)
172      [((Binding.empty, atts), [(thesis, [])])] int
173    |-> Proof.refine_insert
174  end;
175
176in
177
178val consider = gen_consider cert_obtains;
179val consider_cmd = gen_consider read_obtains;
180
181end;
182
183
184
185(** obtain: augmented context based on generalized existence rule **)
186
187(*
188  obtain (a) x where "A x" <proof> \<equiv>
189
190  have thesis if a [intro?]: "\<And>x. A x \<Longrightarrow> thesis" for thesis
191    apply (insert that)
192    <proof>
193  fix x assm <<obtain_export>> "A x"
194*)
195
196local
197
198fun gen_obtain prep_stmt prep_att that_binding raw_decls raw_fixes raw_prems raw_concls int state =
199  let
200    val _ = Proof.assert_forward_or_chain state;
201
202    val ((_, thesis), thesis_ctxt) = obtain_thesis (Proof.context_of state);
203
204    val ({vars, propss, binds, result_binds, ...}, params_ctxt) =
205      prep_stmt (raw_decls @ raw_fixes) (raw_prems @ map #2 raw_concls) thesis_ctxt;
206    val (decls, fixes) = chop (length raw_decls) vars ||> map #2;
207    val (premss, conclss) = chop (length raw_prems) propss;
208    val propss' = (map o map) (Logic.close_prop fixes (flat premss)) conclss;
209
210    val that_prop =
211      Logic.list_rename_params (map (#1 o #2) decls)
212        (fold_rev (Logic.all o #2 o #2) decls (Logic.list_implies (flat propss', thesis)));
213
214    val cparams = map (Thm.cterm_of params_ctxt o #2 o #2) decls;
215    val asms =
216      map (fn ((b, raw_atts), _) => (b, map (prep_att params_ctxt) raw_atts)) raw_concls ~~
217      map (map (rpair [])) propss';
218
219    fun after_qed (result_ctxt, results) state' =
220      let val [rule] = Proof_Context.export result_ctxt (Proof.context_of state') (flat results) in
221        state'
222        |> Proof.fix (map #1 decls)
223        |> Proof.map_context (fold (Variable.bind_term o apsnd (Logic.close_term fixes)) binds)
224        |> Proof.assm (obtain_export params_ctxt rule cparams) [] [] asms
225      end;
226  in
227    state
228    |> Proof.have true NONE after_qed
229      [(Binding.name Auto_Bind.thesisN, NONE, NoSyn)]
230      [((that_binding, [Context_Rules.intro_query NONE]), [(that_prop, [])])]
231      [(Binding.empty_atts, [(thesis, [])])] int
232    |-> Proof.refine_insert
233    |> Proof.map_context (fold Variable.bind_term result_binds)
234  end;
235
236in
237
238val obtain = gen_obtain Proof_Context.cert_stmt (K I);
239val obtain_cmd = gen_obtain Proof_Context.read_stmt Attrib.attribute_cmd;
240
241end;
242
243
244
245(** tactical result **)
246
247fun check_result ctxt thesis th =
248  (case Thm.prems_of th of
249    [prem] =>
250      if Thm.concl_of th aconv thesis andalso
251        Logic.strip_assums_concl prem aconv thesis then th
252      else error ("Guessed a different clause:\n" ^ Thm.string_of_thm ctxt th)
253  | [] => error "Goal solved -- nothing guessed"
254  | _ => error ("Guess split into several cases:\n" ^ Thm.string_of_thm ctxt th));
255
256fun result tac facts ctxt =
257  let
258    val ((thesis_var, thesis), thesis_ctxt) = obtain_thesis ctxt;
259    val st = Goal.init (Thm.cterm_of ctxt thesis);
260    val rule =
261      (case SINGLE (Method.insert_tac thesis_ctxt facts 1 THEN tac thesis_ctxt) st of
262        NONE => raise THM ("Obtain.result: tactic failed", 0, facts)
263      | SOME th =>
264          check_result thesis_ctxt thesis (Raw_Simplifier.norm_hhf thesis_ctxt (Goal.conclude th)));
265
266    val closed_rule = Thm.forall_intr (Thm.cterm_of ctxt (Free thesis_var)) rule;
267    val ((_, [rule']), ctxt') = Variable.import false [closed_rule] ctxt;
268    val obtain_rule =
269      Thm.forall_elim (Thm.cterm_of ctxt (Logic.varify_global (Free thesis_var))) rule';
270    val ((params, stmt), fix_ctxt) = Variable.focus_cterm NONE (Thm.cprem_of obtain_rule 1) ctxt';
271    val (prems, ctxt'') =
272      Assumption.add_assms (obtain_export fix_ctxt obtain_rule (map #2 params))
273        (Drule.strip_imp_prems stmt) fix_ctxt;
274  in ((params, prems), ctxt'') end;
275
276
277
278(** guess: obtain based on tactical result **)
279
280(*
281  <chain_facts>
282  guess x <proof body> <proof end> \<equiv>
283
284  {
285    fix thesis
286    <chain_facts> have "PROP ?guess"
287      apply magic      \<comment> \<open>turn goal into \<open>thesis \<Longrightarrow> #thesis\<close>\<close>
288      <proof body>
289      apply_end magic  \<comment> \<open>turn final \<open>(\<And>x. P x \<Longrightarrow> thesis) \<Longrightarrow> #thesis\<close> into\<close>
290        \<comment> \<open>\<open>#((\<And>x. A x \<Longrightarrow> thesis) \<Longrightarrow> thesis)\<close> which is a finished goal state\<close>
291      <proof end>
292  }
293  fix x assm <<obtain_export>> "A x"
294*)
295
296local
297
298fun unify_params vars thesis_var raw_rule ctxt =
299  let
300    val thy = Proof_Context.theory_of ctxt;
301    val string_of_term = Syntax.string_of_term (Config.put show_types true ctxt);
302
303    fun err msg th = error (msg ^ ":\n" ^ Thm.string_of_thm ctxt th);
304
305    val maxidx = fold (Term.maxidx_typ o snd o fst) vars ~1;
306    val rule = Thm.incr_indexes (maxidx + 1) raw_rule;
307
308    val params = Rule_Cases.strip_params (Logic.nth_prem (1, Thm.prop_of rule));
309    val m = length vars;
310    val n = length params;
311    val _ = m <= n orelse err "More variables than parameters in obtained rule" rule;
312
313    fun unify ((x, T), (y, U)) (tyenv, max) = Sign.typ_unify thy (T, U) (tyenv, max)
314      handle Type.TUNIFY =>
315        err ("Failed to unify variable " ^
316          string_of_term (Free (x, Envir.norm_type tyenv T)) ^ " against parameter " ^
317          string_of_term (Syntax_Trans.mark_bound_abs (y, Envir.norm_type tyenv U)) ^ " in") rule;
318    val (tyenv, _) = fold unify (map #1 vars ~~ take m params)
319      (Vartab.empty, Int.max (maxidx, Thm.maxidx_of rule));
320    val norm_type = Envir.norm_type tyenv;
321
322    val xs = map (apsnd norm_type o fst) vars;
323    val ys = map (apsnd norm_type) (drop m params);
324    val ys' = map Name.internal (Name.variant_list (map fst xs) (map fst ys)) ~~ map #2 ys;
325    val terms = map (Drule.mk_term o Thm.cterm_of ctxt o Free) (xs @ ys');
326
327    val instT =
328      fold (Term.add_tvarsT o #2) params []
329      |> map (fn v => (v, Thm.ctyp_of ctxt (norm_type (TVar v))));
330    val closed_rule = rule
331      |> Thm.forall_intr (Thm.cterm_of ctxt (Free thesis_var))
332      |> Thm.instantiate (instT, []);
333
334    val ((_, rule' :: terms'), ctxt') = Variable.import false (closed_rule :: terms) ctxt;
335    val vars' =
336      map (dest_Free o Thm.term_of o Drule.dest_term) terms' ~~
337      (map snd vars @ replicate (length ys) NoSyn);
338    val rule'' = Thm.forall_elim (Thm.cterm_of ctxt' (Logic.varify_global (Free thesis_var))) rule';
339  in ((vars', rule''), ctxt') end;
340
341fun inferred_type (binding, _, mx) ctxt =
342  let
343    val x = Variable.check_name binding;
344    val ((_, T), ctxt') = Proof_Context.inferred_param x ctxt
345  in ((x, T, mx), ctxt') end;
346
347fun polymorphic ctxt vars =
348  let val Ts = map Logic.dest_type (Variable.polymorphic ctxt (map (Logic.mk_type o #2) vars))
349  in map2 (fn (x, _, mx) => fn T => ((x, T), mx)) vars Ts end;
350
351fun gen_guess prep_var raw_vars int state =
352  let
353    val _ = Proof.assert_forward_or_chain state;
354    val ctxt = Proof.context_of state;
355    val chain_facts = if can Proof.assert_chain state then Proof.the_facts state else [];
356
357    val (thesis_var, thesis) = #1 (obtain_thesis ctxt);
358    val vars = ctxt
359      |> fold_map prep_var raw_vars |-> fold_map inferred_type
360      |> fst |> polymorphic ctxt;
361
362    fun guess_context raw_rule state' =
363      let
364        val ((parms, rule), ctxt') =
365          unify_params vars thesis_var raw_rule (Proof.context_of state');
366        val (xs, _) = Variable.add_fixes (map (#1 o #1) parms) ctxt';
367        val ps = xs ~~ map (#2 o #1) parms;
368        val ts = map Free ps;
369        val asms =
370          Logic.strip_assums_hyp (Logic.nth_prem (1, Thm.prop_of rule))
371          |> map (fn asm => (Term.betapplys (fold_rev Term.abs ps asm, ts), []));
372        val _ = not (null asms) orelse error "Trivial result -- nothing guessed";
373      in
374        state'
375        |> Proof.map_context (K ctxt')
376        |> Proof.fix (map (fn ((x, T), mx) => (Binding.name x, SOME T, mx)) parms)
377        |> `Proof.context_of |-> (fn fix_ctxt => Proof.assm
378          (obtain_export fix_ctxt rule (map (Thm.cterm_of ctxt) ts))
379            [] [] [(Binding.empty_atts, asms)])
380        |> Proof.map_context (fold Variable.unbind_term Auto_Bind.no_facts)
381      end;
382
383    val goal = Var (("guess", 0), propT);
384    val pos = Position.thread_data ();
385    fun print_result ctxt' (k, [(s, [_, th])]) =
386      Proof_Display.print_results int pos ctxt' (k, [(s, [th])]);
387    val before_qed =
388      Method.primitive_text (fn ctxt =>
389        Goal.conclude #> Raw_Simplifier.norm_hhf ctxt #>
390          (fn th => Goal.protect 0 (Conjunction.intr (Drule.mk_term (Thm.cprop_of th)) th)));
391    fun after_qed (result_ctxt, results) state' =
392      let val [_, res] = Proof_Context.export result_ctxt (Proof.context_of state') (flat results)
393      in
394        state'
395        |> Proof.end_block
396        |> guess_context (check_result ctxt thesis res)
397      end;
398  in
399    state
400    |> Proof.enter_forward
401    |> Proof.begin_block
402    |> Proof.fix [(Binding.name Auto_Bind.thesisN, NONE, NoSyn)]
403    |> Proof.chain_facts chain_facts
404    |> Proof.internal_goal print_result Proof_Context.mode_schematic true "guess"
405      (SOME before_qed) after_qed
406      [] [] [(Binding.empty_atts, [(Logic.mk_term goal, []), (goal, [])])]
407    |> snd
408    |> Proof.refine_singleton
409        (Method.primitive_text (fn _ => fn _ => Goal.init (Thm.cterm_of ctxt thesis)))
410  end;
411
412in
413
414val guess = gen_guess Proof_Context.cert_var;
415val guess_cmd = gen_guess Proof_Context.read_var;
416
417end;
418
419end;
420