wenzelm@48985: (* Title: Doc/antiquote_setup.ML wenzelm@21375: Author: Makarius wenzelm@21375: wenzelm@26742: Auxiliary antiquotations for the Isabelle manuals. wenzelm@21375: *) wenzelm@21375: wenzelm@56059: structure Antiquote_Setup: sig end = wenzelm@26742: struct wenzelm@21375: wenzelm@26742: (* misc utils *) wenzelm@26742: wenzelm@29736: fun translate f = Symbol.explode #> map f #> implode; wenzelm@29736: wenzelm@29736: val clean_string = translate wenzelm@26853: (fn "_" => "\\_" wenzelm@30120: | "#" => "\\#" wenzelm@52408: | "$" => "\\$" wenzelm@52408: | "%" => "\\%" wenzelm@26897: | "<" => "$<$" wenzelm@26768: | ">" => "$>$" wenzelm@26768: | "{" => "\\{" wenzelm@30120: | "|" => "$\\mid$" wenzelm@26768: | "}" => "\\}" wenzelm@42666: | "\" => "-" wenzelm@26768: | c => c); wenzelm@26751: wenzelm@31546: fun clean_name "\" = "dots" wenzelm@26897: | clean_name ".." = "ddot" wenzelm@26897: | clean_name "." = "dot" wenzelm@26910: | clean_name "_" = "underscore" wenzelm@26897: | clean_name "{" = "braceleft" wenzelm@26897: | clean_name "}" = "braceright" wenzelm@42666: | clean_name s = s |> translate (fn "_" => "-" | "\" => "-" | c => c); wenzelm@26897: wenzelm@26742: wenzelm@26742: (* ML text *) wenzelm@26742: wenzelm@26742: local wenzelm@26742: wenzelm@59067: fun ml_val (toks1, []) = ML_Lex.read "fn _ => (" @ toks1 @ ML_Lex.read ");" wenzelm@59067: | ml_val (toks1, toks2) = wenzelm@59067: ML_Lex.read "fn _ => (" @ toks1 @ ML_Lex.read " : " @ toks2 @ ML_Lex.read ");"; wenzelm@21375: wenzelm@59067: fun ml_op (toks1, []) = ML_Lex.read "fn _ => (op " @ toks1 @ ML_Lex.read ");" wenzelm@59067: | ml_op (toks1, toks2) = wenzelm@59067: ML_Lex.read "fn _ => (op " @ toks1 @ ML_Lex.read " : " @ toks2 @ ML_Lex.read ");"; wenzelm@46261: wenzelm@59067: fun ml_type (toks1, []) = ML_Lex.read "val _ = NONE : (" @ toks1 @ ML_Lex.read ") option;" wenzelm@55831: | ml_type (toks1, toks2) = wenzelm@59067: ML_Lex.read "val _ = [NONE : (" @ toks1 @ ML_Lex.read ") option, NONE : (" @ wenzelm@59067: toks2 @ ML_Lex.read ") option];"; haftmann@22289: wenzelm@59067: fun ml_exception (toks1, []) = ML_Lex.read "fn _ => (" @ toks1 @ ML_Lex.read " : exn);" wenzelm@55837: | ml_exception (toks1, toks2) = wenzelm@59067: ML_Lex.read "fn _ => (" @ toks1 @ ML_Lex.read " : " @ toks2 @ ML_Lex.read " -> exn);"; haftmann@22289: wenzelm@55831: fun ml_structure (toks, _) = wenzelm@59067: ML_Lex.read "functor XXX() = struct structure XX = " @ toks @ ML_Lex.read " end;"; wenzelm@21375: wenzelm@55831: fun ml_functor (Antiquote.Text tok :: _, _) = wenzelm@59067: ML_Lex.read "ML_Env.check_functor " @ wenzelm@59067: ML_Lex.read (ML_Syntax.print_string (ML_Lex.content_of tok)) wenzelm@55831: | ml_functor _ = raise Fail "Bad ML functor specification"; wenzelm@21375: wenzelm@59082: val is_name = wenzelm@59082: ML_Lex.kind_of #> (fn kind => kind = ML_Lex.Ident orelse kind = ML_Lex.Long_Ident); wenzelm@39869: wenzelm@39869: fun ml_name txt = wenzelm@39869: (case filter is_name (ML_Lex.tokenize txt) of wenzelm@39869: toks as [_] => ML_Lex.flatten toks wenzelm@39869: | _ => error ("Single ML name expected in input: " ^ quote txt)); wenzelm@39869: wenzelm@55831: fun prep_ml source = wenzelm@59066: (Input.source_content source, ML_Lex.read_source false source); wenzelm@55831: wenzelm@37216: fun index_ml name kind ml = Thy_Output.antiquotation name wenzelm@58069: (Scan.lift (Args.text_source_position -- Scan.option (Args.colon |-- Args.text_source_position))) wenzelm@55831: (fn {context = ctxt, ...} => fn (source1, opt_source2) => wenzelm@30394: let wenzelm@55831: val (txt1, toks1) = prep_ml source1; wenzelm@55831: val (txt2, toks2) = wenzelm@55831: (case opt_source2 of wenzelm@55831: SOME source => prep_ml source wenzelm@55831: | NONE => ("", [])); wenzelm@55831: wenzelm@30394: val txt = wenzelm@30394: if txt2 = "" then txt1 wenzelm@30394: else if kind = "type" then txt1 ^ " = " ^ txt2 wenzelm@30394: else if kind = "exception" then txt1 ^ " of " ^ txt2 wenzelm@50239: else if Symbol_Pos.is_identifier (Long_Name.base_name (ml_name txt1)) wenzelm@42290: then txt1 ^ ": " ^ txt2 wenzelm@39858: else txt1 ^ " : " ^ txt2; wenzelm@30394: val txt' = if kind = "" then txt else kind ^ " " ^ txt; wenzelm@55831: wenzelm@59064: val pos = Input.pos_of source1; wenzelm@56275: val _ = wenzelm@57477: ML_Context.eval_in (SOME ctxt) ML_Compiler.flags pos (ml (toks1, toks2)) wenzelm@57477: handle ERROR msg => error (msg ^ Position.here pos); wenzelm@30394: val kind' = if kind = "" then "ML" else "ML " ^ kind; wenzelm@30394: in wenzelm@39869: "\\indexdef{}{" ^ kind' ^ "}{" ^ clean_string (ml_name txt1) ^ "}" ^ wenzelm@58716: (Thy_Output.verbatim_text ctxt txt') wenzelm@30394: end); wenzelm@26742: wenzelm@26742: in wenzelm@21375: wenzelm@56059: val _ = wenzelm@56059: Theory.setup wenzelm@56059: (index_ml @{binding index_ML} "" ml_val #> wenzelm@56059: index_ml @{binding index_ML_op} "infix" ml_op #> wenzelm@56059: index_ml @{binding index_ML_type} "type" ml_type #> wenzelm@56059: index_ml @{binding index_ML_exception} "exception" ml_exception #> wenzelm@56059: index_ml @{binding index_ML_structure} "structure" ml_structure #> wenzelm@56059: index_ml @{binding index_ML_functor} "functor" ml_functor); wenzelm@21375: wenzelm@26742: end; wenzelm@21375: berghofe@23846: wenzelm@30394: (* named theorems *) berghofe@23846: wenzelm@56059: val _ = wenzelm@56059: Theory.setup (Thy_Output.antiquotation @{binding named_thms} wenzelm@43564: (Scan.repeat (Attrib.thm -- Scan.lift (Args.parens Args.name))) wenzelm@43564: (fn {context = ctxt, ...} => wenzelm@43564: map (apfst (Thy_Output.pretty_thm ctxt)) wenzelm@43564: #> (if Config.get ctxt Thy_Output.quotes then map (apfst Pretty.quote) else I) wenzelm@43564: #> (if Config.get ctxt Thy_Output.display wenzelm@43564: then wenzelm@43564: map (fn (p, name) => wenzelm@43564: Output.output (Pretty.string_of (Pretty.indent (Config.get ctxt Thy_Output.indent) p)) ^ wenzelm@43564: "\\rulename{" ^ Output.output (Pretty.str_of (Thy_Output.pretty_text ctxt name)) ^ "}") wenzelm@43564: #> space_implode "\\par\\smallskip%\n" wenzelm@43564: #> enclose "\\begin{isabelle}%\n" "%\n\\end{isabelle}" wenzelm@43564: else wenzelm@43564: map (fn (p, name) => wenzelm@43564: Output.output (Pretty.str_of p) ^ wenzelm@43564: "\\rulename{" ^ Output.output (Pretty.str_of (Thy_Output.pretty_text ctxt name)) ^ "}") wenzelm@43564: #> space_implode "\\par\\smallskip%\n" wenzelm@56059: #> enclose "\\isa{" "}"))); wenzelm@26742: wenzelm@26742: wenzelm@30394: (* theory file *) wenzelm@26742: wenzelm@56059: val _ = wenzelm@56059: Theory.setup (Thy_Output.antiquotation @{binding thy_file} (Scan.lift Args.name) wenzelm@43564: (fn {context = ctxt, ...} => wenzelm@56208: fn name => (Resources.check_thy Path.current name; Thy_Output.output ctxt [Pretty.str name]))); wenzelm@56059: wenzelm@56059: wenzelm@56059: (* Isabelle/jEdit elements *) wenzelm@56059: wenzelm@56059: local wenzelm@56059: wenzelm@56059: fun parse_named a (XML.Elem ((b, props), _)) = wenzelm@56059: (case Properties.get props "NAME" of wenzelm@56059: SOME name => if a = b then [name] else [] wenzelm@56059: | NONE => []) wenzelm@56059: | parse_named _ _ = []; wenzelm@56059: wenzelm@57334: val isabelle_jedit_actions = wenzelm@56135: (case XML.parse (File.read @{path "~~/src/Tools/jEdit/src/actions.xml"}) of wenzelm@56059: XML.Elem (("ACTIONS", _), body) => maps (parse_named "ACTION") body wenzelm@56059: | _ => []); wenzelm@56059: wenzelm@57334: val isabelle_jedit_dockables = wenzelm@56135: (case XML.parse (File.read @{path "~~/src/Tools/jEdit/src/dockables.xml"}) of wenzelm@56059: XML.Elem (("DOCKABLES", _), body) => maps (parse_named "DOCKABLE") body wenzelm@56059: | _ => []); wenzelm@56059: wenzelm@57334: val jedit_actions = wenzelm@57334: Lazy.lazy (fn () => wenzelm@57334: (case Isabelle_System.bash_output wenzelm@57334: "unzip -p \"$JEDIT_HOME/dist/jedit.jar\" org/gjt/sp/jedit/actions.xml" of wenzelm@57334: (txt, 0) => wenzelm@57334: (case XML.parse txt of wenzelm@57334: XML.Elem (("ACTIONS", _), body) => maps (parse_named "ACTION") body wenzelm@57334: | _ => []) wenzelm@57334: | (_, rc) => error ("Cannot unzip jedit.jar\nreturn code = " ^ string_of_int rc))); wenzelm@56059: wenzelm@56059: in wenzelm@56059: wenzelm@57334: fun is_action a = wenzelm@57334: member (op =) isabelle_jedit_actions a orelse wenzelm@57334: member (op =) isabelle_jedit_dockables a orelse wenzelm@57334: member (op =) (Lazy.force jedit_actions) a; wenzelm@56059: wenzelm@56059: end; wenzelm@26742: wenzelm@26751: wenzelm@30394: (* Isabelle/Isar entities (with index) *) wenzelm@26751: wenzelm@26751: local wenzelm@26751: wenzelm@26894: fun no_check _ _ = true; wenzelm@26894: wenzelm@58928: fun is_keyword ctxt (name, _) = wenzelm@58931: Keyword.is_keyword (Thy_Header.get_keywords' ctxt) name; wenzelm@58919: wenzelm@56294: fun check_command ctxt (name, pos) = wenzelm@58928: let wenzelm@58928: val thy = Proof_Context.theory_of ctxt; wenzelm@58928: val keywords = Thy_Header.get_keywords thy; wenzelm@58928: in wenzelm@58928: Keyword.is_command keywords name andalso wenzelm@58923: let wenzelm@58923: val markup = wenzelm@58923: Outer_Syntax.scan keywords Position.none name wenzelm@58928: |> maps (Outer_Syntax.command_reports thy) wenzelm@58923: |> map (snd o fst); wenzelm@58923: val _ = Context_Position.reports ctxt (map (pair pos) markup); wenzelm@58923: in true end wenzelm@58923: end; wenzelm@53061: wenzelm@56467: fun check_system_option ctxt (name, pos) = wenzelm@56467: (Context_Position.report ctxt pos (Options.default_markup (name, pos)); true) wenzelm@56467: handle ERROR _ => false; wenzelm@56467: wenzelm@56294: fun check_tool ctxt (name, pos) = wenzelm@53044: let wenzelm@53044: fun tool dir = wenzelm@53044: let val path = Path.append dir (Path.basic name) wenzelm@53044: in if File.exists path then SOME path else NONE end; wenzelm@53044: in wenzelm@53045: (case get_first tool (Path.split (getenv "ISABELLE_TOOLS")) of wenzelm@53044: NONE => false wenzelm@56294: | SOME path => (Context_Position.report ctxt pos (Markup.path (Path.implode path)); true)) wenzelm@53044: end; wenzelm@28237: wenzelm@26751: val arg = enclose "{" "}" o clean_string; wenzelm@26751: wenzelm@56185: fun entity check markup binding index = wenzelm@37216: Thy_Output.antiquotation wenzelm@56185: (binding |> Binding.map_name (fn name => name ^ wenzelm@43564: (case index of NONE => "" | SOME true => "_def" | SOME false => "_ref"))) wenzelm@53044: (Scan.lift (Scan.optional (Args.parens Args.name) "" -- Parse.position Args.name)) wenzelm@53044: (fn {context = ctxt, ...} => fn (logic, (name, pos)) => wenzelm@30394: let wenzelm@56185: val kind = translate (fn "_" => " " | c => c) (Binding.name_of binding); wenzelm@30394: val hyper_name = wenzelm@30394: "{" ^ Long_Name.append kind (Long_Name.append logic (clean_name name)) ^ "}"; wenzelm@30394: val hyper = wenzelm@30394: enclose ("\\hyperlink" ^ hyper_name ^ "{") "}" #> wenzelm@30394: index = SOME true ? enclose ("\\hypertarget" ^ hyper_name ^ "{") "}"; wenzelm@30394: val idx = wenzelm@30394: (case index of wenzelm@30394: NONE => "" wenzelm@30394: | SOME is_def => wenzelm@30394: "\\index" ^ (if is_def then "def" else "ref") ^ arg logic ^ arg kind ^ arg name); wenzelm@30394: in wenzelm@53044: if check ctxt (name, pos) then wenzelm@30394: idx ^ wenzelm@30394: (Output.output name wenzelm@30394: |> (if markup = "" then I else enclose ("\\" ^ markup ^ "{") "}") wenzelm@38767: |> (if Config.get ctxt Thy_Output.quotes then quote else I) wenzelm@38767: |> (if Config.get ctxt Thy_Output.display wenzelm@38767: then enclose "\\begin{isabelle}%\n" "%\n\\end{isabelle}" wenzelm@30394: else hyper o enclose "\\mbox{\\isa{" "}}")) wenzelm@53044: else error ("Bad " ^ kind ^ " " ^ quote name ^ Position.here pos) wenzelm@30394: end); wenzelm@26897: wenzelm@26894: fun entity_antiqs check markup kind = wenzelm@43564: entity check markup kind NONE #> wenzelm@43564: entity check markup kind (SOME true) #> wenzelm@43564: entity check markup kind (SOME false); wenzelm@26751: wenzelm@26751: in wenzelm@26751: wenzelm@56059: val _ = wenzelm@56059: Theory.setup wenzelm@56185: (entity_antiqs no_check "" @{binding syntax} #> wenzelm@56294: entity_antiqs check_command "isacommand" @{binding command} #> wenzelm@58919: entity_antiqs is_keyword "isakeyword" @{binding keyword} #> wenzelm@58919: entity_antiqs is_keyword "isakeyword" @{binding element} #> wenzelm@56185: entity_antiqs (can o Method.check_name) "" @{binding method} #> wenzelm@56185: entity_antiqs (can o Attrib.check_name) "" @{binding attribute} #> wenzelm@56185: entity_antiqs no_check "" @{binding fact} #> wenzelm@56185: entity_antiqs no_check "" @{binding variable} #> wenzelm@56185: entity_antiqs no_check "" @{binding case} #> wenzelm@56185: entity_antiqs (can o Thy_Output.check_command) "" @{binding antiquotation} #> wenzelm@56185: entity_antiqs (can o Thy_Output.check_option) "" @{binding antiquotation_option} #> wenzelm@58716: entity_antiqs no_check "isasystem" @{binding setting} #> wenzelm@58716: entity_antiqs check_system_option "isasystem" @{binding system_option} #> wenzelm@56185: entity_antiqs no_check "" @{binding inference} #> wenzelm@58716: entity_antiqs no_check "isasystem" @{binding executable} #> wenzelm@56294: entity_antiqs check_tool "isatool" @{binding tool} #> wenzelm@56185: entity_antiqs (can o ML_Context.check_antiquotation) "" @{binding ML_antiquotation} #> wenzelm@58716: entity_antiqs (K (is_action o #1)) "isasystem" @{binding action}); wenzelm@26751: wenzelm@26742: end; wenzelm@26751: wenzelm@26751: end;