src/Pure/Syntax/lexicon.ML
changeset 28407 f9db1da584ac
parent 27887 9f3fd48cf673
child 28413 ee73353fb87c
--- a/src/Pure/Syntax/lexicon.ML	Mon Sep 29 14:41:24 2008 +0200
+++ b/src/Pure/Syntax/lexicon.ML	Mon Sep 29 14:41:25 2008 +0200
@@ -56,7 +56,7 @@
   val tvarT: typ
   val terminals: string list
   val is_terminal: string -> bool
-  val report_token: token -> unit
+  val report_token: Proof.context -> token -> unit
   val matching_tokens: token * token -> bool
   val valued_token: token -> bool
   val predef_term: string -> token option
@@ -162,8 +162,8 @@
   | Comment     => Markup.inner_comment
   | EOF         => Markup.none;
 
-fun report_token (Token (kind, _, (pos, _))) =
-  Position.report (token_kind_markup kind) pos;
+fun report_token ctxt (Token (kind, _, (pos, _))) =
+  ContextPosition.report ctxt (token_kind_markup kind) pos;
 
 
 (* matching_tokens *)