(************************************************************************* * Copyright (C) * 2019 The University of Exeter * 2018-2019 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 *************************************************************************) chapter\Term Antiquotations\ text\Terms are represented by "Inner Syntax" parsed by an Earley parser in Isabelle. For historical reasons, \<^emph>\term antiquotations\ are called therefore somewhat misleadingly "Inner Syntax Antiquotations". \ theory TermAntiquotations imports "Isabelle_DOF.Conceptual" begin text\Since the syntax chosen for values of doc-class attributes is HOL-syntax --- requiring a fast read on the ``What's in Main''-documentation, but not additional knowledge on, say, SML --- an own syntax for references to types, terms, theorems, etc. are necessary. These are the ``Inner Syntax Antiquotations'' since they make only sense \emph{inside} the Inner-Syntax of Isabelle/Isar, so inside the \verb+" ... "+ parenthesis. They are the key-mechanism to denote \<^item> Ontological Links, i.e. attributes refering to document classes defined by the ontology \<^item> Ontological F-Links, i.e. attributes referring to formal entities inside Isabelle (such as thm's) This file contains a number of examples resulting from the % @ {theory "Isabelle_DOF-tests.Conceptual"} does not work here --- why ? \<^theory_text>\Conceptual\ - ontology; the emphasis of this presentation is to present the expressivity of ODL on a paradigmatical example. \ text\Voila the content of the Isabelle_DOF environment so far:\ ML\ val {docobj_tab={tab = x, ...},docclass_tab, ISA_transformer_tab,...} = DOF_core.get_data @{context}; Symtab.dest ISA_transformer_tab; \ text\Some sample lemma:\ lemma murks : "Example=Example" by simp text\Example for a meta-attribute of ODL-type @{typ "file"} with an appropriate ISA for the file @{file "TermAntiquotations.thy"}\ (* not working: text*[xcv::F, u="@{file ''InnerSyntaxAntiquotations.thy''}"]\Lorem ipsum ...\ *) text*[xcv1::A, x=5]\Lorem ipsum ...\ text*[xcv3::A, x=7]\Lorem ipsum ...\ text\Example for a meta-attribute of ODL-type @{typ "typ"} with an appropriate ISA for the theorem @{thm "refl"}}\ text*[xcv2::C, g="@{thm ''HOL.refl''}"]\Lorem ipsum ...\ text\Major sample: test-item of doc-class \F\ with a relational link between class instances, and links to formal Isabelle items like \typ\, \term\ and \thm\. \ text*[xcv4::F, r="[@{thm ''HOL.refl''}, @{thm \TermAntiquotations.murks\}]", (* long names required *) b="{(@{docitem ''xcv1''},@{docitem \xcv2\})}", (* notations \...\ vs. ''...'' *) s="[@{typ \int list\}]", properties = "[@{term \H \ H\}]" (* notation \...\ required for UTF8*) ]\Lorem ipsum ...\ text*[xcv5::G, g="@{thm \HOL.sym\}"]\Lorem ipsum ...\ text\... and here we add a relation between @{docitem \xcv3\} and @{docitem \xcv2\} into the relation \verb+b+ of @{docitem \xcv5\}. Note that in the link-relation, a @{typ "C"}-type is required, but a @{typ "G"}-type is offered which is legal in \verb+Isa_DOF+ because of the sub-class relation between those classes: \ update_instance*[xcv4::F, b+="{(@{docitem ''xcv3''},@{docitem ''xcv5''})}"] text\And here is the result on term level:\ ML\ @{docitem_attribute b::xcv4} \ end