chapter \The Document Ontology Common Library for the Isabelle Ontology Framework\ text\ Offering \<^item> ... \<^item> \<^item> LaTeX support. \ theory Isa_COL imports Isa_DOF begin section\ Library of Standard Text Ontology \ datatype placement = pl_h | pl_t | pl_b | pl_ht | pl_hb doc_class figure = relative_width :: "int" (* percent of textwidth *) src :: "string" placement :: placement spawn_columns :: bool <= True doc_class side_by_side_figure = figure + anchor :: "string" caption :: "string" relative_width2 :: "int" (* percent of textwidth *) src2 :: "string" anchor2 :: "string" caption2 :: "string" doc_class figure_group = (* trace :: "doc_class rexp list" <= "[]" automatically generated since monitor clause *) caption :: "string" rejects figure_group (* this forbids recursive figure-groups not supported by the current LaTeX style-file. *) accepts "\figure\\<^sup>+" (* dito the future table *) (* dito the future monitor: table - block *) text\ The attribute @{term "level"} in the subsequent enables doc-notation support section* etc. we follow LaTeX terminology on levels \<^enum> part = Some -1 \<^enum> chapter = Some 0 \<^enum> section = Some 1 \<^enum> subsection = Some 2 \<^enum> subsubsection = Some 3 \<^enum> ... for scholarly paper: invariant level > 0 \ doc_class text_element = level :: "int option" <= "None" referentiable :: bool <= "False" variants :: "String.literal set" <= "{STR ''outline'', STR ''document''}" section\Some attempt to model standardized links to Standard Isabelle Formal Content\ doc_class assertions = properties :: "term list" doc_class "thms" = properties :: "thm list" doc_class formal_item = item :: "(assertions + thms)" doc_class formal_content = style :: "string option" accepts "\formal_item\\<^sup>+" doc_class concept = tag :: "string" <= "''''" properties :: "thm list" <= "[]" section\Tests\ ML\@{term "side_by_side_figure"}; @{typ "doc_class rexp"}; DOF_core.SPY; \ end