# HG changeset patch # User wenzelm # Date 962392871 -7200 # Node ID 2651a4db888354bc64f25aa7ac0d59926d556f51 # Parent 4afe62073b41845b767c870898cf67f3703fbd6d help_antiquotations; diff -r 4afe62073b41 -r 2651a4db8883 src/Pure/Isar/isar_output.ML --- a/src/Pure/Isar/isar_output.ML Fri Jun 30 17:51:56 2000 +0200 +++ b/src/Pure/Isar/isar_output.ML Fri Jun 30 21:21:11 2000 +0200 @@ -10,6 +10,7 @@ sig val add_commands: (string * (Args.src -> Proof.context -> string)) list -> unit val add_options: (string * (string -> (unit -> string) -> unit -> string)) list -> unit + val help_antiquotations: unit -> Pretty.T list val boolean: string -> bool val integer: string -> int val args: (Proof.context * Args.T list -> 'a * (Proof.context * Args.T list)) -> @@ -68,6 +69,11 @@ fun options [] f = f | options (opt :: opts) f = option opt (options opts f); + +fun help_antiquotations () = + [Pretty.strs ("antiquotation commands:" :: Symtab.keys (! global_commands)), + Pretty.strs ("antiquotation options:" :: Symtab.keys (! global_options))]; + end; diff -r 4afe62073b41 -r 2651a4db8883 src/Pure/Isar/outer_syntax.ML --- a/src/Pure/Isar/outer_syntax.ML Fri Jun 30 17:51:56 2000 +0200 +++ b/src/Pure/Isar/outer_syntax.ML Fri Jun 30 21:21:11 2000 +0200 @@ -222,6 +222,7 @@ Toplevel.keep (fn state => let val opt_thy = try Toplevel.theory_of state in help_outer_syntax () @ + IsarOutput.help_antiquotations () @ Method.help_methods opt_thy @ Attrib.help_attributes opt_thy |> Pretty.chunks |> Pretty.writeln