diff --git a/Isabelle_DOF-Unit-Tests/Concept_OntoReferencing.thy b/Isabelle_DOF-Unit-Tests/Concept_OntoReferencing.thy index c2b1f9e1..5a915fbc 100644 --- a/Isabelle_DOF-Unit-Tests/Concept_OntoReferencing.thy +++ b/Isabelle_DOF-Unit-Tests/Concept_OntoReferencing.thy @@ -84,11 +84,12 @@ text*[d::D, a1 = "X3"] \ ... phasellus amet id massa nunc, pede suscipit r text\Not only text-elements are "ontology-aware", proofs and code can this be too !\ references\ \ \Referencing from and to a ML-code context:\ +(*<*) ML*[c4::C, z = "Some @{A \a1\}"]\ fun fac x = if x = 0 then 1 else x * (fac(x-1)) val v = \<^value_>\A.x (the (z @{C \c4\}))\ |> HOLogic.dest_number |> snd |> fac \ - +(*>*) definition*[a2::A, x=5, level="Some 1"] xx' where "xx' \ A.x @{A \a1\}" if "A.x @{A \a1\} = 5" lemma*[e5::E] testtest : "xx + A.x @{A \a1\} = yy + A.x @{A \a1\} \ xx = yy" by simp @@ -140,7 +141,7 @@ update_instance*[f::F,b:="{(@{docitem \a\}::A,@{docitem \c1 section\Closing the Monitor and testing the Results.\ close_monitor*[struct] -text\@{A a0}\ + text\And the trace of the monitor is:\ ML\val trace = @{trace_attribute struct}\ ML\@{assert} (trace = diff --git a/Isabelle_DOF-Unit-Tests/High_Level_Syntax_Invariants.thy b/Isabelle_DOF-Unit-Tests/High_Level_Syntax_Invariants.thy index eb5f373e..8500e310 100644 --- a/Isabelle_DOF-Unit-Tests/High_Level_Syntax_Invariants.thy +++ b/Isabelle_DOF-Unit-Tests/High_Level_Syntax_Invariants.thy @@ -1,11 +1,52 @@ +(************************************************************************* + * Copyright (C) + * 2019-2023 The University of Exeter + * 2018-2023 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\High level syntax Invariants\ theory High_Level_Syntax_Invariants imports "Isabelle_DOF.Isa_DOF" - "Isabelle_DOF-Unit-Tests_document" - + "Isabelle_DOF-Unit-Tests_document" + TestKit begin +section\Test Purpose.\ +text\ + Without invariants, ontological classes as such are too liberal in many situations. + Similarly to UML constraints, invariants or hand-programmed checking functions + can be added in ODL ontologies in order to constrain class instances or + (via monitor traces) impose structural constraints over an entire document. + + While hand-programmed checking functions were tested in test-case + \<^verbatim>\Concept_Example_Low_Level_Invariant\, in this text case, we test + high-level invariants, i.e. data-constraints speicified as executable + HOL-predicates in the @{theory_text \invariant\} clause of ODL definitions. + + To enable the checking of the invariants, the \invariants_checking\ + theory attribute must be set:\ + + +section\The Scenario.\ + +text\ This is merely an example that shows that the generated invariants + fit nicely together; i.e. allow for sensible consistency and invariant + preservation proofs related to ontological matchings. \ + + +text\Using HOL, we can define a mapping between two ontologies. + It is called ontology matching or ontology alignment. + Here is an example which show how to map two classes. + HOL also allows us to map the invariants (ontological rules) of the classes!\ + text\ Ontological classes as described so far are too liberal in many situations. There is a first high-level syntax implementation for class invariants.