1 (* Title: Pure/Isar/obtain.ML
3 Author: Markus Wenzel, TU Muenchen
4 License: GPL (GNU GENERAL PUBLIC LICENSE)
6 The 'obtain' language element -- generalized existence at the level of
10 obtain x where "P x" <proof> ==
14 assume that [intro]: "!!x. P x ==> thesis"
15 <chain_facts> have thesis <proof (insert that)>
17 fix x assm (obtained) "P x"
23 val obtain: (string list * string option) list ->
24 ((string * Proof.context attribute list) * (string * (string list * string list)) list) list
25 -> Proof.state -> Proof.state Seq.seq
26 val obtain_i: (string list * typ option) list ->
27 ((string * Proof.context attribute list) * (term * (term list * term list)) list) list
28 -> Proof.state -> Proof.state Seq.seq
31 structure Obtain: OBTAIN =
37 fun export_obtain state parms rule _ cprops thm =
39 val {sign, prop, maxidx, ...} = Thm.rep_thm thm;
40 val cparms = map (Thm.cterm_of sign) parms;
43 |> Drule.implies_intr_goals cprops
44 |> Drule.forall_intr_list cparms
45 |> Drule.forall_elim_vars (maxidx + 1);
46 val elim_tacs = replicate (length cprops) (Tactic.etac Drule.triv_goal);
48 val concl = Logic.strip_assums_concl prop;
49 val bads = parms inter (Term.term_frees concl);
51 if not (null bads) then
52 raise Proof.STATE ("Conclusion contains obtained parameters: " ^
53 space_implode " " (map (ProofContext.string_of_term (Proof.context_of state)) bads), state)
54 else if not (ObjectLogic.is_judgment sign (Logic.strip_assums_concl prop)) then
55 raise Proof.STATE ("Conclusions of 'obtain' context must be object-logic judgments", state)
56 else (Tactic.rtac thm' THEN' RANGE elim_tacs) 1 rule
65 fun gen_obtain prep_vars prep_propp raw_vars raw_asms state =
67 val _ = Proof.assert_forward_or_chain state;
68 val chain_facts = if Proof.is_chain state then Proof.the_facts state else [];
69 val thy = Proof.theory_of state;
70 val sign = Theory.sign_of thy;
73 val (vars_ctxt, vars) = foldl_map prep_vars (Proof.context_of state, raw_vars);
74 val xs = flat (map fst vars);
75 val fix_ctxt = vars_ctxt |> ProofContext.fix_i vars;
78 val (asms_ctxt, proppss) = prep_propp (fix_ctxt, map snd raw_asms);
79 val asm_props = flat (map (map fst) proppss);
80 val asms = map fst raw_asms ~~ proppss;
82 val _ = ProofContext.warn_extra_tfrees fix_ctxt asms_ctxt;
85 val thesisN = Term.variant xs AutoBind.thesisN;
87 ProofContext.bind_skolem fix_ctxt [thesisN] (ObjectLogic.fixed_judgment sign thesisN);
89 fun occs_var x = Library.get_first (fn t =>
90 ProofContext.find_free t (ProofContext.get_skolem fix_ctxt x)) asm_props;
91 val raw_parms = map occs_var xs;
92 val parms = mapfilter I raw_parms;
94 mapfilter (fn (Some (Free a), x) => Some (a, x) | _ => None) (raw_parms ~~ xs);
97 Term.list_all_free (map #1 parm_names, Logic.list_implies (asm_props, bound_thesis))
98 |> Library.curry Logic.list_rename_params (map #2 parm_names);
100 fun after_qed st = st
102 |> Seq.map (fn st' => st'
104 |> Proof.assm_i (export_obtain state parms (Proof.the_fact st')) asms);
107 |> Proof.enter_forward
109 |> Proof.fix_i [([thesisN], None)]
110 |> Proof.assume_i [((thatN, [Method.intro_local]), [(that_prop, ([], []))])]
113 |> Proof.from_facts chain_facts
114 |> Proof.have_i after_qed "" [] (bound_thesis, ([], []))
115 |> Method.refine (Method.Basic (K (Method.insert (Proof.the_facts state')))))
118 val obtain = gen_obtain ProofContext.read_vars ProofContext.read_propp;
119 val obtain_i = gen_obtain ProofContext.cert_vars ProofContext.cert_propp;