(************************************************************************* * Copyright (C) * 2019-2021 The University of Exeter * 2018-2021 The 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 "02_Background" imports "01_Introduction" begin (*>*) chapter*[background::text_section]\ Background\ section*[bgrnd1::introduction]\The Isabelle System Architecture\ figure*[architecture::figure,relative_width="95",src="''figures/isabelle-architecture''"]\ The system architecture of Isabelle (left-hand side) and the asynchronous communication between the Isabelle system and the IDE (right-hand side). \ text*[bg::introduction]\ While Isabelle @{cite "nipkow.ea:isabelle:2002"} is widely perceived as an interactive theorem prover for HOL (Higher-order Logic)~@{cite "nipkow.ea:isabelle:2002"}, we would like to emphasize the view that Isabelle is far more than that: it is the \<^emph>\Eclipse of Formal Methods Tools\. This refers to the ``\<^emph>\generic system framework of Isabelle/Isar underlying recent versions of Isabelle. Among other things, Isar provides an infrastructure for Isabelle plug-ins, comprising extensible state components and extensible syntax that can be bound to ML programs. Thus, the Isabelle/Isar architecture may be understood as an extension and refinement of the traditional `LCF approach', with explicit infrastructure for building derivative systems.\''~@{cite "wenzel.ea:building:2007"} The current system framework offers moreover the following features: \<^item> a build management grouping components into to pre-compiled sessions, \<^item> a prover IDE (PIDE) framework~@{cite "wenzel:asynchronous:2014"} with various front-ends \<^item> documentation-generation, \<^item> code generators for various target languages, \<^item> an extensible front-end language Isabelle/Isar, and, \<^item> last but not least, an LCF style, generic theorem prover kernel as the most prominent and deeply integrated system component. The Isabelle system architecture shown in @{docitem \architecture\} comes with many layers, with Standard ML (SML) at the bottom layer as implementation language. The architecture actually foresees a \<^emph>\Nano-Kernel\ (our terminology) which resides in the SML structure\<^boxed_sml>\Context\. This structure provides a kind of container called \<^emph>\context\ providing an identity, an ancestor-list as well as typed, user-defined state for components (plugins) such as \<^isadof>. On top of the latter, the LCF-Kernel, tactics, automated proof procedures as well as specific support for higher specification constructs were built.\ section*[dof::introduction]\The Document Model Required by \<^dof>\ text\ In this section, we explain the assumed document model underlying our Document Ontology Framework (\<^dof>) in general. In particular we discuss the concepts \<^emph>\integrated document\\<^bindex>\integrated document\, \<^emph>\sub-document\\<^bindex>\sub-document\, \<^emph>\text-element\\<^bindex>\text-element\, and \<^emph>\semantic macros\\<^bindex>\semantic macros\ occurring inside text-elements. Furthermore, we assume two different levels of parsers (for \<^emph>\outer\ and \<^emph>\inner syntax\) where the inner-syntax is basically a typed \\\-calculus and some Higher-order Logic (HOL)\<^bindex>\HOL\. \ (*<*) declare_reference*["fig:dependency"::text_section] (*>*) text\ We assume a hierarchical document model\<^index>\document model\, \<^ie>, an \<^emph>\integrated\ document consist of a hierarchy \<^emph>\sub-documents\ (files) that can depend acyclically on each other. Sub-documents can have different document types in order to capture documentations consisting of documentation, models, proofs, code of various forms and other technical artifacts. We call the main sub-document type, for historical reasons, \<^emph>\theory\-files. A theory file\<^bindex>\theory!file\ consists of a \<^emph>\header\\<^bindex>\header\, a \<^emph>\context definition\\<^index>\context\, and a body consisting of a sequence of \<^emph>\command\s (see @{figure (unchecked) "fig:dependency"}). Even the header consists of a sequence of commands used for introductory text elements not depending on any context. The context-definition contains an \<^boxed_theory_text>\import\ and a \<^boxed_theory_text>\keyword\ section, for example: @{boxed_theory_text [display]\ theory Example \\Name of the 'theory'\ imports \\Declaration of 'theory' dependencies\ Main \\Imports a library called 'Main'\ keywords \\Registration of keywords defined locally\ requirement \\A command for describing requirements\ \} where \<^boxed_theory_text>\Example\ is the abstract name of the text-file, \<^boxed_theory_text>\Main\ refers to an imported theory (recall that the import relation must be acyclic) and \<^boxed_theory_text>\keywords\ are used to separate commands from each other. \ text\ A text-element \<^index>\text-element\ may look like this: @{boxed_theory_text [display]\ text\ According to the \<^emph>\reflexivity\ axiom @{thm refl}, we obtain in \ for @{term "fac 5"} the result @{value "fac 5"}.\\} so it is a command \<^theory_text>\text\ followed by an argument (here in \\ ... \\ paranthesis) which contains characters and and a special notation for semantic macros \<^bindex>\semantic macros\ (here \<^theory_text>\@{term "fac 5"}).\ \ text\ We distinguish fundamentally two different syntactic levels: \<^item> the \<^emph>\outer-syntax\\<^bindex>\syntax!outer\\<^index>\outer syntax|see {syntax, outer}\ (\<^ie>, the syntax for commands) is processed by a lexer-library and parser combinators built on top, and \<^item> the \<^emph>\inner-syntax\\<^bindex>\syntax!inner\\<^index>\inner syntax|see {syntax, inner}\ (\<^ie>, the syntax for \\\-terms in HOL) with its own parametric polymorphism type checking. On the semantic level, we assume a validation process for an integrated document, where the semantics of a command is a transformation \\ \ \\ for some system state \\\. This document model can be instantiated with outer-syntax commands for common text elements, \<^eg>, \<^theory_text>\section\...\\ or \<^theory_text>\text\...\\. Thus, users can add informal text to a sub-document using a text command: @{boxed_theory_text [display] \text\This is a description.\\ } This will type-set the corresponding text in, for example, a PDF document. However, this translation is not necessarily one-to-one: text elements can be enriched by formal, \<^ie>, machine-checked content via \<^emph>\semantic macros\, called antiquotations\<^bindex>\antiquotation\: @{boxed_theory_text [display] \text\ According to the \<^emph>\reflexivity\ axiom @{thm refl}, we obtain in \ for @{term "fac 5"} the result @{value "fac 5"}.\\ } which is represented in the final document (\<^eg>, a PDF) by: @{boxed_pdf [display] \According to the $\emph{reflexivity}$ axiom $\mathrm{x = x}$, we obtain in $\Gamma$ for $\operatorname{fac} \text{\textrm{5}}$ the result $\text{\textrm{120}}$.\ } Semantic macros are partial functions of type \\ \ text\; since they can use the system state, they can perform all sorts of specific checks or evaluations (type-checks, executions of code-elements, references to text-elements or proven theorems such as \refl\, which is the reference to the axiom of reflexivity). Semantic macros establish \<^emph>\formal content\ inside informal content; they can be type-checked before being displayed and can be used for calculations before being typeset. They represent the device for linking the formal with the informal. \ figure*["fig:dependency"::figure,relative_width="70",src="''figures/document-hierarchy''"] \A Theory-Graph in the Document Model. \ section*[bgrnd21::introduction]\Implementability of the Required Document Model\ text\ Batch-mode checkers for \<^dof> can be implemented in all systems of the LCF-style prover family, \<^ie>, systems with a type-checked \term\, and abstract \thm\-type for theorems (protected by a kernel). This includes, \<^eg>, ProofPower, HOL4, HOL-light, Isabelle, or Coq and its derivatives. \<^dof> is, however, designed for fast interaction in an IDE. If a user wants to benefit from this experience, only Isabelle and Coq have the necessary infrastructure of asynchronous proof-processing and support by a PIDE~@{cite "wenzel:asynchronous:2014" and "wenzel:system:2014" and "barras.ea:pervasive:2013" and "faithfull.ea:coqoon:2018"} which in many features over-accomplishes the required features of \<^dof>. For example, current Isabelle versions offer cascade-syntaxes (different syntaxes and even parser-technologies which can be nested along the \\...\\ barriers), while \<^dof> actually only requires a two-level syntax model. \ figure*["fig:dof-ide"::figure,relative_width="95",src="''figures/cicm2018-combined''"]\ The \<^isadof> IDE (left) and the corresponding PDF (right), showing the first page of~@{cite "brucker.ea:isabelle-ontologies:2018"}.\ text\ We call the present implementation of \<^dof> on the Isabelle platform \<^isadof> . @{docitem "fig:dof-ide"} shows a screen-shot of an introductory paper on \<^isadof>~@{cite "brucker.ea:isabelle-ontologies:2018"}: the \<^isadof> PIDE can be seen on the left, while the generated presentation in PDF is shown on the right. Isabelle provides, beyond the features required for \<^dof>, a lot of additional benefits. Besides UTF8-support for characters used in text-elements, Isabelle offers built-in already a mechanism user-programmable antiquotations \<^index>\antiquotations\ which we use to implement semantic macros \<^index>\semantic macros\ in \<^isadof> (We will actually use these two terms as synonym in the context of \<^isadof>). Moreover, \<^isadof> allows for the asynchronous evaluation and checking of the document content~@{cite "wenzel:asynchronous:2014" and "wenzel:system:2014" and "barras.ea:pervasive:2013"} and is dynamically extensible. Its PIDE provides a \<^emph>\continuous build, continuous check\ functionality, syntax highlighting, and auto-completion. It also provides infrastructure for displaying meta-information (\<^eg>, binding and type annotation) as pop-ups, while hovering over sub-expressions. A fine-grained dependency analysis allows the processing of individual parts of theory files asynchronously, allowing Isabelle to interactively process large (hundreds of theory files) documents. Isabelle can group sub-documents into sessions, \<^ie>, sub-graphs of the document-structure that can be ``pre-compiled'' and loaded instantaneously, \<^ie>, without re-processing, which is an important means to scale up. \ (*<*) end (*>*)