src/Pure/Isar/isar_thy.ML
changeset 8450 dc44d6533f0f
parent 8428 be4c8a57cf7e
child 8466 f7b06595d0c8
--- a/src/Pure/Isar/isar_thy.ML	Tue Mar 14 11:32:38 2000 +0100
+++ b/src/Pure/Isar/isar_thy.ML	Tue Mar 14 11:33:14 2000 +0100
@@ -79,7 +79,9 @@
   val fix_i: ((string list * typ option) * Comment.text) list -> ProofHistory.T -> ProofHistory.T
   val match_bind: ((string list * string) * Comment.text) list -> ProofHistory.T -> ProofHistory.T
   val match_bind_i: ((term list * term) * Comment.text) list -> ProofHistory.T -> ProofHistory.T
-  val invoke_case: string * Comment.text -> ProofHistory.T -> ProofHistory.T
+  val invoke_case: (string * Args.src list) * Comment.text -> ProofHistory.T -> ProofHistory.T
+  val invoke_case_i: (string * Proof.context attribute list) * Comment.text
+    -> ProofHistory.T -> ProofHistory.T
   val theorem: (bstring * Args.src list * (string * (string list * string list)))
     * Comment.text -> bool -> theory -> ProofHistory.T
   val theorem_i: (bstring * theory attribute list * (term * (term list * term list)))
@@ -288,7 +290,10 @@
 val fix_i = ProofHistory.apply o Proof.fix_i o map Comment.ignore;
 val match_bind = ProofHistory.apply o Proof.match_bind o map Comment.ignore;
 val match_bind_i = ProofHistory.apply o Proof.match_bind_i o map Comment.ignore;
-val invoke_case = ProofHistory.apply o Proof.invoke_case o Comment.ignore;
+
+fun invoke_case ((name, src), comment_ignore) = ProofHistory.apply (fn state =>
+  Proof.invoke_case (name, map (Attrib.local_attribute (Proof.theory_of state)) src) state);
+val invoke_case_i = ProofHistory.apply o Proof.invoke_case o Comment.ignore;
 
 
 (* statements *)