diff -r 474263d29057 -r b7b916a82dca src/Pure/Isar/obtain.ML --- a/src/Pure/Isar/obtain.ML Mon Nov 13 21:59:49 2000 +0100 +++ b/src/Pure/Isar/obtain.ML Mon Nov 13 22:01:07 2000 +0100 @@ -21,10 +21,10 @@ signature OBTAIN = sig val obtain: ((string list * string option) * Comment.text) list - * ((string * Args.src list * (string * (string list * string list)) list) + * (((string * Args.src list) * (string * (string list * string list)) list) * Comment.text) list -> ProofHistory.T -> ProofHistory.T val obtain_i: ((string list * typ option) * Comment.text) list - * ((string * Proof.context attribute list * (term * (term list * term list)) list) + * (((string * Proof.context attribute list) * (term * (term list * term list)) list) * Comment.text) list -> ProofHistory.T -> ProofHistory.T end; @@ -66,6 +66,7 @@ let val _ = Proof.assert_forward_or_chain state; val chain_facts = if Proof.is_chain state then Proof.the_facts state else []; + val thy = Proof.theory_of state; (*obtain vars*) val (vars_ctxt, vars) = @@ -78,15 +79,12 @@ (bind_skolem prop, (map bind_skolem pats1, map bind_skolem pats2)); (*obtain asms*) - fun prep_asm (ctxt, (name, src, raw_propps)) = - let - val atts = map (prep_att (ProofContext.theory_of ctxt)) src; - val (ctxt', propps) = foldl_map prep_propp (ctxt, raw_propps); - in (ctxt', ((name, atts, map bind_propp propps), map #1 propps)) end; + val (asms_ctxt, proppss) = prep_propp (vars_ctxt, map (snd o Comment.ignore) raw_asms); + val asm_props = flat (map (map fst) proppss); - val (asms_ctxt, asms_result) = foldl_map prep_asm (vars_ctxt, map Comment.ignore raw_asms); - val (asms, asm_propss) = Library.split_list asms_result; - val asm_props = flat asm_propss; + fun prep_asm ((name, src), propps) = ((name, map (prep_att thy) src), map bind_propp propps); + val asms = map2 prep_asm (map (fst o Comment.ignore) raw_asms, proppss); + val _ = ProofContext.warn_extra_tfrees vars_ctxt asms_ctxt; (*that_prop*) @@ -96,7 +94,7 @@ val xs' = mapfilter occs_var xs; val parms = map (bind_skolem o Free) xs'; - val bound_thesis = bind_skolem (AutoBind.atomic_judgment (Proof.theory_of state) thesisN); + val bound_thesis = bind_skolem (AutoBind.atomic_judgment thy thesisN); val that_prop = Term.list_all_free (xs', Logic.list_implies (asm_props, bound_thesis)); fun export_obtained rule = @@ -112,7 +110,7 @@ |> Proof.enter_forward |> Proof.begin_block |> Proof.fix_i [([thesisN], None)] - |> Proof.assume_i [(thatN, [Method.intro_local], [(that_prop, ([], []))])] + |> Proof.assume_i [((thatN, [Method.intro_local]), [(that_prop, ([], []))])] |> (fn state' => state' |> Proof.from_facts chain_facts @@ -139,7 +137,7 @@ (Scan.optional (P.and_list1 (Scan.repeat1 P.name -- Scan.option (P.$$$ "::" |-- P.typ) -- P.marg_comment) --| P.$$$ "where") [] -- - P.and_list1 ((P.opt_thm_name ":" -- Scan.repeat1 P.propp >> P.triple1) -- P.marg_comment) + P.and_list1 (P.opt_thm_name ":" -- Scan.repeat1 P.propp -- P.marg_comment) >> (Toplevel.print oo (Toplevel.proof o obtain))); val _ = OuterSyntax.add_keywords ["where"];