Theory M_05_Proofs_Ontologies
theory
"M_05_Proofs_Ontologies"
imports
"M_04_Document_Ontology"
begin
chapter*[onto_proofs::technical]‹Proofs over Ontologies›
text‹It is a distinguishing feature of \<^dof> that it does not directly generate meta-data rather
than generating a ∗‹theory of meta-data› that can be used in HOL-terms on various levels
of the Isabelle-system and its document generation technology. Meta-data theories can be
converted into executable code and efficiently used in validations, but also used for theoretic
reasoning over given ontologies. While the full potential of this latter possibility still
needs to be explored, we present in the following sections two applications:
▸ Verified ontological morphisms, also called ∗‹ontological mappings› in the literature
\<^cite>‹"books/daglib/0032976" and "atl" and "BGPP95"›, \<^ie> proofs of invariant preservation
along translation-functions of all instances of ▩‹doc_class›-es.
▸ Verified refinement relations between the structural descriptions of theory documents,
\<^ie> proofs of language inclusions of monitors of global ontology classes.
›
section*["morphisms"::scholarly_paper.text_section] ‹Proving Properties over Ontologies›
subsection‹Ontology-Morphisms: a Prototypical Example›
text‹We define a small ontology with the following classes:›
doc_class AA = aa :: nat
doc_class BB = bb :: int
doc_class CC = cc :: int
doc_class DD = dd :: int
doc_class EE = ee :: int
doc_class FF = ff :: int
onto_morphism (AA, BB) to CC and (DD, EE) to FF
where "convert⇩A⇩A⇩×⇩B⇩B⇩⇒⇩C⇩C σ = ⦇ CC.tag_attribute = 1::int,
CC.cc = int(aa (fst σ)) + bb (snd σ)⦈"
and "convert⇩D⇩D⇩×⇩E⇩E⇩⇒⇩F⇩F σ = ⦇ FF.tag_attribute = 1::int,
FF.ff = dd (fst σ) + ee (snd σ) ⦈"
text‹Note that the \<^term>‹convert⇩A⇩A⇩×⇩B⇩B⇩⇒⇩C⇩C›-morphism involves a data-format conversion, and that the
resulting transformation of @{doc_class AA}-instances and @{doc_class BB}-instances is surjective
but not injective. The \<^term>‹CC.tag_attribute› is used to potentially differentiate instances with
equal attribute-content and is irrelevant here.›
doc_class A_A = aa :: nat
doc_class BB' = bb :: int
onto_morphism (A_A, BB', CC, DD, EE) to FF
where "convert⇩A⇩_⇩A⇩×⇩B⇩B⇩'⇩×⇩C⇩C⇩×⇩D⇩D⇩×⇩E⇩E⇩⇒⇩F⇩F σ = ⦇ FF.tag_attribute = 1::int,
FF.ff = int(aa (fst σ)) + bb (fst (snd σ))⦈"
text‹This specification construct introduces the following constants and definitions:
▪ @{term [source] ‹convert⇩A⇩A⇩_⇩B⇩B⇩_⇩C⇩C :: AA × BB ⇒ CC›}
▪ @{term [source] ‹convert⇩D⇩D⇩_⇩E⇩E⇩_⇩F⇩F :: DD × EE ⇒ FF›}
% @{term [source] ‹convert⇩A⇩_⇩A⇩×⇩B⇩B⇩'⇩×⇩C⇩C⇩×⇩D⇩D⇩×⇩E⇩E⇩⇒⇩F⇩F :: A_A × BB' × CC × DD × EE ⇒ FF›}
and corresponding definitions. ›
subsection‹Proving the Preservation of Ontological Mappings : A Document-Ontology Morphism›
text‹\<^dof> as a system is currently particularly geared towards ∗‹document›-ontologies, in
particular for documentations generated from Isabelle theories. We used it meanwhile for the
generation of various conference and journal papers, notably using the
\<^theory>‹Isabelle_DOF.scholarly_paper› and \<^theory>‹Isabelle_DOF.technical_report›-ontologies,
targeting a (small) variety of \<^LaTeX> style-files. A particular aspect of these ontologies,
especially when targeting journals from publishers such as ACM, Springer or Elsevier, is the
description of publication meta-data. Publishers tend to have their own styles on what kind
meta-data should be associated with a journal publication; thus, the depth and
precise format of affiliations, institution, their relation to authors, and author descriptions
(with photos or without, hair left-combed or right-combed, etcpp...) varies.
In the following, we present an attempt to generalized ontology with several ontology mappings
to more specialized ones such as concrete journals and/or the \<^theory>‹Isabelle_DOF.scholarly_paper›-
ontology which we mostly used for our own publications.
›
doc_class elsevier_org =
organization :: string
address_line :: string
postcode :: int
city :: string