(************************************************************************* * Copyright (C) * 2019-2022 University of Exeter * 2018-2022 University of Paris-Saclay * 2018 The University of Sheffield * * License: * This program can be redistributed and/or modified under the terms * of the 2-clause BSD-style license. * * SPDX-License-Identifier: BSD-2-Clause *************************************************************************) (*<*) theory "M_04_RefMan" imports "M_03_GuidedTour" begin declare_reference*[infrastructure::technical] (*>*) chapter*[isadof_ontologies::technical]\Ontologies and their Development\ text\ In this chapter, we explain the concepts of \<^isadof> in a more systematic way, and give guidelines for modeling new ontologies, present underlying concepts for a mapping to a representation, and give hints for the development of new document templates. \<^isadof> is embedded in the underlying generic document model of Isabelle as described in \<^introduction>\dof\. Recall that the document language can be extended dynamically, \<^ie>, new \<^emph>\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 provides a number of new command definitions in Isabelle's document model. \<^isadof> consists consists basically of five components: \<^item> the \<^emph>\core\ in \<^theory>\Isabelle_DOF.Isa_DOF\ providing the \<^emph>\ontology definition language\ (ODL) which allow for the definitions of document-classes and necessary datatypes, \<^item> the \<^emph>\core\ also provides an own \<^emph>\family of commands\ such as \<^boxed_theory_text>\text*\, \<^boxed_theory_text>\ML*\, \<^boxed_theory_text>\value*\ , \<^etc>.; They allow for the annotation of document-elements with meta-information defined in ODL, \<^item> the \<^isadof> library \<^theory>\Isabelle_DOF.Isa_COL\ providing ontological basic (documents) concepts \<^bindex>\ontology\ as well as supporting infrastructure, \<^item> an infrastructure for ontology-specific \<^emph>\layout definitions\, exploiting this meta-information, and \<^item> an infrastructure for generic \<^emph>\layout definitions\ for documents following, \<^eg>, the format guidelines of publishers or standardization bodies. \ text\ Similarly to Isabelle, which is based on a core logic \<^theory>\Pure\ and then extended by libraries to major systems like \<^verbatim>\HOL\, \<^isadof> has a generic core infrastructure \<^dof> and then presents itself to users via major library extensions, which add domain-specific system-extensions. Ontologies\<^bindex>\ontology\ in \<^isadof> are not just a sequence of descriptions in \<^isadof>'s Ontology Definition Language (ODL)\<^bindex>\ODL\. Rather, they are themselves presented as integrated sources that provide textual descriptions, abbreviations, macro-support and even ML-code. Conceptually, the library of \<^isadof> is currently organized as follows\<^footnote>\The \<^emph>\technical\ organization is slightly different and shown in @{technical (unchecked) \infrastructure\}.\ \<^bindex>\COL\\<^bindex>\scholarly\_paper\\<^bindex>\technical\_report\ \<^bindex>\CENELEC\: % \<^latex>\ \begin{center} \begin{minipage}{.9\textwidth}\footnotesize \dirtree{% .1 COL\DTcomment{The Common Ontology Library}. .2 scholarly\_paper\DTcomment{Scientific Papers}. .3 technical\_report\DTcomment{Extended Papers, Technical Reports}. .4 CENELEC\_50128\DTcomment{Papers according to CENELEC\_50128}. .4 CC\_v3\_1\_R5\DTcomment{Papers according to Common Criteria}. .4 \ldots. } \end{minipage} \end{center}\ These libraries not only provide ontological concepts, but also syntactic sugar in Isabelle's command language Isar that is of major importance for users (and may be felt as \<^isadof> key features by many authors). In reality, they are derived concepts from more generic ones; for example, the commands \<^boxed_theory_text>\title*\, \<^boxed_theory_text>\section*\, \<^boxed_theory_text>\subsection*\, \<^etc>, are in reality a kind of macros for \<^boxed_theory_text>\text*[