(*<*) theory "04_RefMan" imports "03_GuidedTour" (* apparently inaccessible; for railroads "Isar_Ref.Base" *) begin (*>*) chapter*[isadof::technical,main_author="Some(@{docitem ''bu''}::author)"] \ \isadof : Syntax and Semantics\ text\ \isadof is embedded in the underlying generic document model of Isabelle as described in @{docitem "sec:background"}. Recall that the document language can be extended dynamically, \ie, new \user-defined\ can be introduced at run-time. This is similar to the definition of new functions in an interpreter. \isadof as a system plugin is is a number of new command definitions in Isabelle's document model. \isadof consists consists basically of three components: \<^item> an own \<^emph>\family of text-element\ such as @{command "title*"}, @{command "chapter*"} @{command "text*"}, etc., which can be annotated with meta-information defined in the underlying ontology definition and allow to build a \<^emph>\core\ document, \<^item> the \<^emph>\ontology definition\ which is an Isabelle theory file with definitions for document-classes and all auxiliary datatypes (we call this the Ontology Definition Language (ODL)) \<^item> the \<^emph>\layout definition\ for the given ontology exploiting this meta-information. \ text\ Note that the document core \<^emph>\may\, but \<^emph>\must\ not use Isabelle definitions or proofs for checking the formal content---the present paper is actually an example of a document not containing any proof. Consequently, the document editing and checking facility provided by \isadof addresses the needs of common users for an advanced text-editing environment, neither modeling nor proof knowledge is inherently required. We expect authors of ontologies to have experience in the use of \isadof, basic modeling (and evtl. SML programming) experience and, last but not least, domain knowledge of the ontology to be modeled. Users with experience in UML-like meta-modeling will feel familiar with most concepts; however, we expect no need for insight in the Isabelle proof language, for example, or other more advanced concepts. The document generation process of \isadof is currently restricted to \LaTeX, which means that the layout is defined by a set of \LaTeX{} style files. Several layout definitions for one ontology are possible which paves the way for different \<^emph>\views\ on the same integrated document, addressing the needs of different purposes and/or target readers. Conceiving new style files will definitively require knowledge over \LaTeX{} and some knowledge over the Isabelle document generation process. The situation is roughly similar to \LaTeX{}-users, who usually have minimal knowledge about the content in style-files (\<^verbatim>\.sty\-files). In the document core authors \<^emph>\can\ use \LaTeX{} commands in their source, but this limits the possibility of using different representation technologies, \eg, HTML, and increases the risk of arcane error-messages in the generated \LaTeX{}. Using low-level \LaTeX{} commands is at the users risk. A correctly checked \isadof document should typeset, provided that a few basic pitfalls are avoided: no dangling \<^verbatim>\{\ or \<^verbatim>\}\, no spontaneous unprotected \<^verbatim>\_\, etc. It is also helpful to execute the final @{command "check_doc_global"} command to check the global reference stucture. \ section\Ontology Modeling in ODL\ text\As already mentioned, ODL has some similarities to meta-modeling languages such as UML class models: It builds upon concepts like class, inheritance, class-instances, attributes, references to instances, and class-invariants. Some more advanced concepts like advanced type-checking, referencing to formal entities of Isabelle, and monitors are due to its specific application in the Isabelle context. Conceptually, ontologies specified in ODL consist of: \<^item> \<^emph>\document classes\ (syntactically marked by the \inlineisar{doc_class} keyword) that describe concepts; \<^item> an optional document base class expressing single inheritance extensions; \<^item> \<^emph>\attributes\ specific to document classes, where \<^item> attributes are typed; \<^item> attributes of instances of document elements are mutable; \<^item> attributes can refer to other document classes, thus, document classes must also be HOL-types (such attributes are called \<^emph>\links\); \<^item> a special link, the reference to a super-class, establishes an \<^emph>\is-a\ relation between classes; \<^item> classes may refer to other classes via a regular expression in a \<^emph>\where\ clause (classes with a where clauses are called \<^emph>\monitor classes\); \<^item> attributes may have default values in order to facilitate notation. \ text\ The \isadof ontology specification language consists basically on a notation for document classes, where the attributes were typed with HOL-types and can be instantiated by terms HOL-terms, \ie, the actual parsers and type-checkers of the Isabelle system were reused. This has the particular advantage that \isadof commands can be arbitrarily mixed with Isabelle/HOL commands providing the machinery for type declarations and term specifications such as enumerations. In particular, document class definitions provide: \<^item> a HOL-type for each document class as well as inheritance, \<^item> support for attributes with HOL-types and optional default values, \<^item> support for overriding of attribute defaults but not overloading, and \<^item> text-elements annotated with document classes; they are mutable instances of document classes. \ text\ Attributes referring to other ontological concepts are called \<^emph>\links\. The HOL-types inside the document specification language support built-in types for Isabelle/HOL \inlineisar+typ+'s, \inlineisar+term+'s, and \inlineisar+thm+'s reflecting internal Isabelle's internal types for these entities; when denoted in HOL-terms to instantiate an attribute, for example, there is a specific syntax (called \<^emph>\inner syntax antiquotations\) that is checked by \isadof for consistency. Document classes can have a \inlineisar+where+ clause containing a regular expression over class names. Classes with such a \inlineisar+where+ were called \<^emph>\monitor classes\. While document classes and their inheritance relation structure meta-data of text-elements in an object-oriented manner, monitor classes enforce structural organization of documents via the language specified by the regular expression enforcing a sequence of text-elements that must belong to the corresponding classes. \ text\ A major design decision of ODL is to denote attribute values by HOL-terms and HOL-types. Consequently, ODL can refer to any predefined type defined in the HOL library, \eg, \inlineisar+string+ or \inlineisar+int+ as well as parameterized types, \eg, % \inlineisar+_ option+, \inlineisar+_ list+, \inlineisar+_ set+, or products \inlineisar+_ \ _+. As a consequence of the document model, ODL definitions may be arbitrarily intertwined with standard HOL type definitions. Finally, document class definitions result in themselves in a HOL-types in order to allow \<^emph>\links\ to and between ontological concepts. \ text\ \ subsection\Annotating with Ontology Meta-Data: Outer Syntax\ text\\dof introduces its own family of text-commands, which allows having side effects of the global context \inlineisar+\+ and thus to store and manage own meta-information. Among others, \dof provides the commands \inlineisar+section*[]\...\+, \inlineisar+subsection*[]\...\Cclose>+, or \inlineisar+text*[]\...\+. Here, the argument \inlineisar++ is a syntax for declaring instance, class and attributes for this text element, following the scheme \begin{isar} :: , attr_1 = , ..., attr_n = \end{isar} The \inlineisar++ can be omitted, which represents the implicit superclass \inlineisar+text+, where \inlineisar+attr_i+ must be declared attributes in the class and where the HOL \inlineisar++ must have the corresponding HOL type. Attributes from a class definition may be left undefined; definitions of attribute values \<^emph>\override\ default values or values of super-classes. Overloading of attributes is not permitted in \dof. \ text\ We can now annotate a text as follows. First, we have to place a particular document into the context of our conceptual example ontology (\autoref{lst:doc}): \begin{isar} theory Steam_Boiler imports tiny_cert (* The ontology defined in Listing 1.1. *) begin \end{isar} This opens a new document (theory), called \texttt{Steam\_Boiler} that imports our conceptual example ontology ``\texttt{tiny\_cert}'' (stored in a file \texttt{tiny\_cert.thy}).\footnote{The usual import-mechanisms of the Isabelle document model applies also to ODL: ontologies can be extended, several ontologies may be imported, a document can validate several ontologies.} \noindent Now we can continue to annotate our text as follows: \begin{isar} title*[a] \The Steam Boiler Controller\ abstract*[abs, safety_level="SIL4", keywordlist = "[''controller'']"]\ We present a formalization of a program which serves to control the level of water in a steam boiler. \ section*[intro::introduction]\Introduction\ text\We present ...\ section*[T1::technical]\Physical Environment\ text\ The system comprises the following units [*] the steam-boiler [*] a device to measure the quantity of water in the steam-boiler [*] ... \ \end{isar} \ text\ Where \inlineisar+title*[a ...]+ is a predefined macro for \inlineisar+text*[a::title,...]\...\+ (similarly \inlineisar+abstract*+). The macro \inlineisar+section*+ assumes a class-id referring to a class that has a \inlineisar+level+ attribute. We continue our example text: \begin{isar} text*[c1::contrib_claim, based_on="[''pumps'',''steam boiler'']" ]\ As indicated in @{introduction "intro"}, we the water level of the boiler is always between the minimum and the maximum allowed level. \ \end{isar} \ text\ The first text element in this example fragment \<^emph>\defines\ the text entity \inlineisar+c1+ and also references the formerly defined text element \inlineisar+intro+ (which will be represented in the PDF output, for example, by a text anchor ``Section 1'' and a hyperlink to its beginning). The antiquotation \inlineisar+{introduction ...}+, which is automatically generated from the ontology, is immediately validated (the link to \inlineisar+intro+ is defined) and type-checked (it is indeed a link to an \inlineisar+introduction+ text-element). Moreover, the IDE automatically provides editing and development support such as auto-completion or the possibility to ``jump'' to its definition by clicking on the antiquotation. The consistency checking ensures, among others, that the final document will not contain any ``dangling references'' or references to entities of another type. \ text\ \dof as such does not require a particular evaluation strategy; however, if the underlying implementation is based on a declaration-before-use strategy, a mechanism for forward declarations of references is necessary: \begin{isar} declare_reference* [] \end{isar} This command declares the existence of a text-element and allows for referencing it, although the actual text-element will occur later in the document. \ text \ \subsection{Editing Documents with Ontology Meta-Data: Inner Syntax} We continue our running example as follows: \begin{isar}[mathescape] text*[d1::definition]\ We define the water level {term "level"} of a system state {term "\"} of the steam boiler as follows: \ definition level :: "state \ real" where "level \ = level0 + ..." update_instance*[d1::definition, property += "[{term ''level \ = level0 + ...''}]"] text*[only::result,evidence="proof"]\ The water level is never lower than {term "level0"}: \ theorem level_always_above_level_0: "\ \. level \ \ level0" unfolding level_def by auto update_instance*[only::result, property += "[{thm ''level_always_above_level_0''}]"] \end{isar} \ text\ As mentioned earlier, instances of document classes are mutable. We use this feature to modify meta-data of these text-elements and ``assign'' them to the property-list afterwards and add results from Isabelle definitions and proofs. The notation \inlineisar|A+=X| stands for \inlineisar|A := A + X|. This mechanism can also be used to define the required relation between \<^emph>\claims\ and \<^emph>\results\ required in the \inlineisar|establish|-relation required in a \inlineisar|summary|. \ section*["odl-manual"::technical]\The ODL Manual\ subsection*["odl-manual0"::technical]\The Isabelle/HOL Specification Language\ text\The \isadof ontology definition language (ODL) is an extension of Isabelle/HOL; document class definitions can therefore be arbitrarily mixed with standard HOL specification constructs. In order to spare the user of ODL a lenghty analysis of @{cite "isaisarrefman19" and "datarefman19" and "functions19"}, we present syntax and semantics of the specification constructs that are most likely relevant for the developer of ontologies. Note that our presentation is actually a simplification of the original sources following the needs of our target audience here; for a full- blown account, the reader is referred to the original descriptions.\ text\ \<^item> \name\ : with the syntactic category of \name\'s we refer to alpha-numerical identifiers (called \short_id\'s in @{cite "isaisarrefman19"}) and identifiers in \inlineisar+" ... "+ which might additionally contain certian ``quasi-letters'' such as \inlineisar+_+, \inlineisar+-+, \inlineisar+.+, etc. Details in @{cite "isaisarrefman19"}. \<^item> \tyargs\ : \<^rail>\ typefree | ('(' (typefree * ',') ')')\ \typefree\ denotes fixed type variable(\'a\, \'b\, ...) (c.f. @{cite "isaisarrefman19"}) \<^item> \dt_name\ : \<^rail>\ (tyargs?) name (mixfix?)\ The syntactic entity \name\ denotes an identifier, \mixfix\ denotes the usual parenthesized mixfix notation (see @{cite "isaisarrefman19"}). The \name\'s referred here are type names such as \<^verbatim>\int\, \<^verbatim>\string\, \<^verbatim>\list\, \<^verbatim>\set\, etc. \<^item> \type_spec\ : \<^rail>\ (tyargs?) name\ The \name\'s referred here are type names such as \<^verbatim>\int\, \<^verbatim>\string\, \<^verbatim>\list\, \<^verbatim>\set\, etc. \<^item> \type\ : \<^rail>\ (( '(' ( type * ',') ')' )? name) | typefree \ \<^item> \dt_ctor\ : \<^rail>\ name (type*) (mixfix?)\ \<^item> \datatype_specification\ : \<^rail>\ @@{command "datatype"} dt_name '=' (dt_ctor * '|' )\ \<^item> \type_synonym_specification\ : \<^rail>\ @@{command "type_synonym"} type_spec '=' type\ \<^item> \constant_definition\ : \<^rail>\ @@{command "definition"} name '::' type 'where' '"' name '=' expr '"'\ \<^item> \expr\ : the syntactic category \expr\ here denotes the very rich ``inner-syntax'' language of mathematical notations for $\lambda$-terms in Isabelle/HOL. Example expressions are: \inlineisar|1+2| (arithmetics), \inlineisar|[1,2,3]| (lists), \inlineisar|''ab c''| (strings), \inlineisar|{1,2,3}| (sets), \inlineisar|(1,2,3)| (tuples), \inlineisar|\ x. P(x) \ Q x = C| (formulas), etc. For a comprehensive overview, consult the summary ``What's in Main'' @{cite "nipkowMain19"}. \ text\Other specification constructs, which are potentially relevant for an (advanced) ontology developer, might be recursive function definitions with pattern-matching @{cite "functions19"}, extensible record specifications, @{cite "isaisarrefman19"} and abstract type declarations.\ subsection*["odl-manual1"::technical]\The Command Syntax for Document Class Specifications\ text\ \<^item> \class_id\ : \class_id\'s are type-\name\'s that have been introduced via a \doc_class_specification\. \<^item> \doc_class_specification\ : \<^rail>\ @@{command "doc_class"} class_id '=' (class_id '+')? (attribute_decl+) \ (accepts_clause rejects_clause?)?\ document classes whose specification contains an \accepts_clause\ are called \<^emph>\monitor classes\ or \<^emph>\monitors\ for short. \<^item> \attribute_decl\ : \<^rail>\ name '::' '"' type '"' default_clause? \ \<^item> \accepts_clause\ : \<^rail>\ 'accepts' '"' regexpr '"'\ \<^item> \rejects_clause\ : \<^rail>\ 'rejects' (class_id * ',') \ \<^item> \default_clause\ : \<^rail>\ '<=' '"' expr '"' \ \<^item> \regexpr\ : \<^rail>\ '\' class_id '\' | '(' regexpr ')' | (regexpr '||' regexpr) | (regexpr '~~' regexpr) | ('\' regexpr '\') | ( '\' regexpr '\\<^sup>*') \ regular expressions describe sequences of \class_id\s (and indirect sequences of document items corresponding to the \class_id\s). The constructors for alternative, sequence, repetitions and non-empty sequence follow in the top-down order of the above diagram. \ section*["odl-design"::technical]\The Design of ODL\ text\ We illustrate the design of \dof by modeling a small ontology that can be used for writing formal specifications that, \eg, could build the basis for an ontology for certification documents used in processes such as Common Criteria~@{cite "cc:cc-part3:2006"} or CENELEC 50128~@{cite "bsi:50128:2014"}.@{footnote \The \isadof distribution contains an ontology for writing documents for a certification according to CENELEC 50128.\} Moreover, in examples of certification documents, we refer to a controller of a steam boiler that is inspired by the famous steam boiler formalization challenge~@{cite "abrial:steam-boiler:1996"}. \ text\ \begin{isar}[float,mathescape,label={lst:doc},caption={An example ontology modeling simple certification documents, including scientific papers such as~@{cite "brucker.ea:isabelle-ontologies:2018"}; also recall \autoref{fig:dof-ide}.}] doc_class title = short_title :: "string option" <= "None" doc_class author = email :: "string" <= "''''" datatype classification = SIL0 | SIL1 | SIL2 | SIL3 | SIL4 doc_class abstract = keywordlist :: "string list" <= [] safety_level :: "classification" <= "SIL3" doc_class text_section = authored_by :: "author set" <= "{}" level :: "int option" <= "None" type_synonym notion = string doc_class introduction = text_section + authored_by :: "author set" <= "UNIV" uses :: "notion set" doc_class claim = introduction + based_on :: "notion list" doc_class technical = text_section + formal_results :: "thm list" doc_class "d$$efinition" = technical + is_formal :: "bool" property :: "term list" <= "[]" datatype kind = expert_opinion | argument | proof doc_class result = technical + evidence :: kind property :: "thm list" <= "[]" doc_class example = technical + referring_to :: "(notion + d$$efinition) set" <= "{}" doc_class "conclusion" = text_section + establish :: "(claim \ result) set" \end{isar}%$ \ text\ \autoref{lst:doc} shows an example ontology for mathematical papers (an extended version of this ontology was used for writing @{cite "brucker.ea:isabelle-ontologies:2018"}, also recall \autoref{fig:dof-ide}). The commands \inlineisar+datatype+ (modeling fixed enumerations) and \inlineisar+type_synonym+ (defining type synonyms) are standard mechanisms in HOL systems. Since ODL is an add-on, we have to quote sometimes constant symbols (\eg, \inlineisar+"proof"+) to avoid confusion with predefined keywords. ODL admits overriding (such as \inlineisar+authored_by+ in the document class \inlineisar+introduction+), where it is set to another library constant, but no overloading. All \inlineisar+text_section+ elements have an optional \inlineisar+level+ attribute, which will be used in the output generation for the decision if the context is a section header and its level (\eg, chapter, section, subsection). While \within\ an inheritance hierarchy overloading is prohibited, attributes may be re-declared freely in independent parts (as is the case for \inlineisar+property+). \ (* subsection*[onto_future::technical]\ Monitor Classes \ text\ Besides sub-typing, there is another relation between document classes: a class can be a \<^emph>\monitor\ to other ones, which is expressed by the occurrence of a \inlineisar+where+ clause in the document class definition containing a regular expression (see @{example (unchecked) \scholar_onto\}). While class-extension refers to data-inheritance of attributes, a monitor imposes structural constraints -- the order -- in which instances of monitored classes may occur. \ text\ The control of monitors is done by the commands: \<^item> \inlineisar+open_monitor* + \<^item> \inlineisar+close_monitor* + \ text\ where the automaton of the monitor class is expected to be in a final state. In the final state, user-defined SML Monitors can be nested, so it is possible to "overlay" one or more monitoring classes and imposing different sets of structural constraints in a Classes which are neither directly nor indirectly (via inheritance) mentioned in the monitor are \<^emph>\independent\ from a monitor; instances of independent test elements may occur freely. \ *) subsection*["sec:monitors"::technical]\ODL Monitors\ text\ We call a document class with an accept-clause a \<^emph>\monitor\. Syntactically, an accept-clause contains a regular expression over class identifiers. We can extend our \inlineisar+tiny_cert+ ontology with the following example: \begin{isar} doc_class article = style_id :: string <= "''CENELEC50128''" accepts "(title ~~ \author\\+\ ~~ abstract ~~ \introduction\\+\ ~~ \technical || example\\+\ ~~ \conclusion\\+\)" \end{isar} Semantically, monitors introduce a behavioral element into ODL: \begin{isar} open_monitor*[this::article] (* begin of scope of monitor "this" *) ... close_monitor*[this] (* end of scope of monitor "this" *) \end{isar} Inside the scope of a monitor, all instances of classes mentioned in its accept-clause (the \<^emph>\accept-set\) have to appear in the order specified by the regular expression; instances not covered by an accept-set may freely occur. Monitors may additionally contain a reject-clause with a list of class-ids (the reject-list). This allows specifying ranges of admissible instances along the class hierarchy: \begin{compactitem} \item a superclass in the reject-list and a subclass in the accept-expression forbids instances superior to the subclass, and \item a subclass $S$ in the reject-list and a superclass $T$ in the accept-list allows instances of superclasses of $T$ to occur freely, instances of $T$ to occur in the specified order and forbids instances of $S$. \end{compactitem} Monitored document sections can be nested and overlap; thus, it is possible to combine the effect of different monitors. For example, it would be possible to refine the \inlineisar+example+ section by its own monitor and enforce a particular structure in the presentation of examples. Monitors manage an implicit attribute \inlineisar+trace+ containing the list of ``observed'' text element instances belonging to the accept-set. Together with the concept of ODL class invariants, it is possible to specify properties of a sequence of instances occurring in the document section. For example, it is possible to express that in the sub-list of \inlineisar+introduction+-elements, the first has an \inlineisar+introduction+ element with a \inlineisar+level+ strictly smaller than the others. Thus, an introduction is forced to have a header delimiting the borders of its representation. Class invariants on monitors allow for specifying structural properties on document sections. \ subsection*["odl-manual2"::technical]\Examples\ subsection\Meta-types as Types\ text\To express the dependencies between text elements to the formal entities, \eg, \inlinesml+term+ ($\lambda$-term), \inlinesml+typ+, or \inlinesml+thm+, we represent the types of the implementation language \<^emph>\inside\ the HOL type system. We do, however, not reflect the data of these types. They are just declared abstract types, ``inhabited'' by special constant symbols carrying strings, for example of the format \inlineisar+{thm }+. When HOL expressions were used to denote values of \inlineisar+doc_class+ instance attributes, this requires additional checks after conventional type-checking that this string represents actually a defined entity in the context of the system state \inlineisar+\+. For example, the \inlineisar+establish+ attribute in the previous section is the power of the ODL: here, we model a relation between \<^emph>\claims\ and \<^emph>\results\ which may be a formal, machine-check theorem of type \inlinesml+thm+ denoted by, for example: \inlineisar+property = "[{thm ''system_is_safe''}]"+ in a system context \inlineisar+\+ where this theorem is established. Similarly, attribute values like \inlineisar+property = "{term \A \ B\}"+ require that the HOL-string \inlineisar+A \ B+ is again type-checked and represents indeed a formula in $\theta$. Another instance of this process, which we call \second-level type-checking\, are term-constants generated from the ontology such as \inlineisar+{definition }+. For the latter, the argument string must be checked that it represents a reference to a text-element having the type \inlineisar+definition+ according to the ontology in \autoref{lst:doc}.\ subsection*["sec:class_inv"::technical]\ODL Class Invariants\ text\ Ontological classes as described so far are too liberal in many situations. For example, one would like to express that any instance of a \inlineisar+result+ class finally has a non-empty property list, if its \inlineisar+kind+ is \inlineisar+proof+, or that the \inlineisar+establish+ relation between \inlineisar+claim+ and \inlineisar+result+ is surjective. In a high-level syntax, this type of constraints could be expressed, \eg, by: \begin{isar} \ x \ result. x@kind = proof \ x@kind \ [] \ x \ conclusion. \ y \ Domain(x@establish) \ \ y \ Range(x@establish). (y,z) \ x@establish \ x \ introduction. finite(x@authored_by) \end{isar} @{cartouche [display=true] \ (* 1 *) \ x \ result. x@kind = proof \ x@kind \ [] (* 2 *) \ x \ conclusion. \ y \ Domain(x@establish) \ \ y\ Range(x@establish). (y,z) \ x@establish (* 3 *) \ x \ introduction. finite(x@authored_by) \} \fixme{experiment... } %% where \inlineisar+result+, \inlineisar+conclusion+, and %% where \result\, \conclusion\, and \introduction\ are the set of all possible instances of these document classes. All specified constraints are already checked in the IDE of \dof while editing; it is however possible to delay a final error message till the closing of a monitor (see next section). The third constraint enforces that the user sets the \authored_by\ set, otherwise an error will be %%user sets the \inlineisar+authored_by+ set, otherwise an error will be reported. \ section*["core-manual"::technical]\Annotatable Text-Elements for Core-Documents\ text\In general, all standard text-elements from the Isabelle document model such as @{command "chapter"}, @{command "section"}, @{command "text"}, have in the \isadof implementation their counterparts in the family of text-elements that are "ontology-aware", \ie they dispose on a meta-argument list that allows to define that a test-element \<^enum> has an identity as a text-object labelled as \obj_id\ \<^enum> belongs to a document class \class_id\ that has been defined earlier \<^enum> has its class-attributes set with particular values (which are denotable in Isabelle/HOL mathematical term syntax) \ subsection*["core-manual1"::technical]\Syntax\ text\The syntax of toplevel \isadof commands reads as follows: \<^item> \annotated_text_element\ : \<^rail>\ ( ( @@{command "title*"} | @@{command "subtitle*"} | @@{command "chapter*"} | @@{command "section*"} | @@{command "subsection*"} | @@{command "subsubsection*"} | @@{command "paragraph*"} | @@{command "subparagraph*"} | @@{command "text*"} | @@{command "figure*"} | @@{command "side_by_side_figure*"} | @@{command "open_monitor*"} | @@{command "close_monitor*"} | @@{command "Definition*"} | @@{command "Lemma*"} | @@{command "Theorem*"} | @@{command "Conjecture*"} ) '[' meta_args ']' '\' text '\' ) | change_status_command | inspection_command \ \<^item> \meta_args\ : \<^rail>\(obj_id ('::' class_id) ((attribute '=' term)) * ',')\ \<^item> \rich_meta_args\ : \<^rail>\ (obj_id ('::' class_id) ((attribute (('=' | '+=') term)) * ','))\ \<^item> \isadof\change_status_command\ : \<^rail>\ (@@{command "update_instance*"} '[' rich_meta_args ']') | (@@{command "declare_reference*"} (obj_id ('::' class_id)))\ \<^item> \isadof\inspection_command\ : \<^rail>\ @@{command "print_doc_classes"} | @@{command "print_doc_items"} | @@{command "check_doc_global"}\ \ subsection*["commonlib"::technical]\Common Ontology Library (COL)\ subsection*["core-manual2"::technical]\Examples\ (*<*) end (*>*)