src/Pure/Pure.thy
changeset 78792 103467dc5117
parent 78740 45ff003d337c
child 78796 f34926a91fea
equal deleted inserted replaced
78791:4f7dce5c1a81 78792:103467dc5117
     6 
     6 
     7 theory Pure
     7 theory Pure
     8 keywords
     8 keywords
     9     "!!" "!" "+" ":" ";" "<" "<=" "==" "=>" "?" "[" "\<comment>" "\<equiv>" "\<leftharpoondown>" "\<rightharpoonup>" "\<rightleftharpoons>"
     9     "!!" "!" "+" ":" ";" "<" "<=" "==" "=>" "?" "[" "\<comment>" "\<equiv>" "\<leftharpoondown>" "\<rightharpoonup>" "\<rightleftharpoons>"
    10     "\<subseteq>" "]" "binder" "by" "in" "infix" "infixl" "infixr" "is" "open" "output"
    10     "\<subseteq>" "]" "binder" "by" "in" "infix" "infixl" "infixr" "is" "open" "output"
    11     "overloaded" "pervasive" "premises" "structure" "unchecked"
    11     "overloaded" "passive" "pervasive" "premises" "structure" "unchecked"
    12   and "private" "qualified" :: before_command
    12   and "private" "qualified" :: before_command
    13   and "assumes" "constrains" "defines" "fixes" "for" "if" "includes" "notes" "rewrites"
    13   and "assumes" "constrains" "defines" "fixes" "for" "if" "includes" "notes" "rewrites"
    14     "obtains" "shows" "when" "where" "|" :: quasi_command
    14     "obtains" "shows" "when" "where" "|" :: quasi_command
    15   and "text" "txt" :: document_body
    15   and "text" "txt" :: document_body
    16   and "text_raw" :: document_raw
    16   and "text_raw" :: document_raw
   332       >> (fn (pervasive, txt) => Isar_Cmd.declaration {syntax = true, pervasive = pervasive} txt));
   332       >> (fn (pervasive, txt) => Isar_Cmd.declaration {syntax = true, pervasive = pervasive} txt));
   333 
   333 
   334 val _ =
   334 val _ =
   335   Outer_Syntax.local_theory \<^command_keyword>\<open>simproc_setup\<close> "define simproc in ML"
   335   Outer_Syntax.local_theory \<^command_keyword>\<open>simproc_setup\<close> "define simproc in ML"
   336     (Parse.name_position --
   336     (Parse.name_position --
   337       (\<^keyword>\<open>(\<close> |-- Parse.enum1 "|" Parse.term --| \<^keyword>\<open>)\<close> --| \<^keyword>\<open>=\<close>) --
   337       (\<^keyword>\<open>(\<close> |-- Parse.enum1 "|" Parse.term --| \<^keyword>\<open>)\<close>) --
   338       Parse.ML_source >> (fn ((a, b), c) => Isar_Cmd.simproc_setup a b c));
   338       (\<^keyword>\<open>=\<close> |-- Parse.ML_source) --
       
   339       Parse.opt_keyword "passive"
       
   340     >> (fn (((a, b), c), d) => Isar_Cmd.simproc_setup a b c d));
   339 
   341 
   340 in end\<close>
   342 in end\<close>
   341 
   343 
   342 
   344 
   343 subsection \<open>Theory commands\<close>
   345 subsection \<open>Theory commands\<close>