src/Pure/Pure.thy
changeset 62913 13252110a6fe
parent 62902 3c0f53eae166
child 62944 3ee643c5ed00
--- a/src/Pure/Pure.thy	Thu Apr 07 22:09:23 2016 +0200
+++ b/src/Pure/Pure.thy	Fri Apr 08 20:15:20 2016 +0200
@@ -9,7 +9,7 @@
     "!!" "!" "+" "--" ":" ";" "<" "<=" "=" "=>" "?" "[" "\<comment>" "\<equiv>"
     "\<leftharpoondown>" "\<rightharpoonup>" "\<rightleftharpoons>"
     "\<subseteq>" "]" "assumes" "attach" "binder" "constrains"
-    "defines" "rewrites" "fixes" "for" "identifier" "if" "in" "includes" "infix"
+    "defines" "rewrites" "fixes" "for" "if" "in" "includes" "infix"
     "infixl" "infixr" "is" "notes" "obtains" "open" "output"
     "overloaded" "pervasive" "premises" "private" "qualified" "rewrites"
     "shows" "structure" "unchecked" "where" "when" "|"
@@ -210,8 +210,7 @@
   Outer_Syntax.local_theory @{command_keyword simproc_setup} "define simproc in ML"
     (Parse.position Parse.name --
       (@{keyword "("} |-- Parse.enum1 "|" Parse.term --| @{keyword ")"} --| @{keyword "="}) --
-      Parse.ML_source -- Scan.optional (@{keyword "identifier"} |-- Scan.repeat1 Parse.xname) []
-    >> (fn (((a, b), c), d) => Isar_Cmd.simproc_setup a b c d));
+      Parse.ML_source >> (fn ((a, b), c) => Isar_Cmd.simproc_setup a b c));
 
 in end\<close>