(************************************************************************* * 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\A conceptual introduction into DOF and its features:\ theory Conceptual imports "Isabelle_DOF.Isa_DOF" "Isabelle_DOF.Isa_COL" begin section\Excursion: On the semantic consequences of this definition: \ text\Consider the following document class definition and its consequences:\ doc_class A = level :: "int option" x :: int text\This class definition leads an implicit Isabelle/HOL \<^theory_text>\record\ definition (cf. \<^url>\https://isabelle.in.tum.de/doc/isar-ref.pdf\, chapter 11.6.). Consequently, \<^theory_text>\doc_class\'es inherit the entire theory-infrastructure from Isabelle records: \<^enum> there is a HOL-type \<^typ>\A\ and its extensible version \<^typ>\'a A_scheme\ \<^enum> there are HOL-terms representing \<^emph>\doc\_class instances\ with the high-level syntax: \<^enum> \<^term>\undefined\level := Some (1::int), x := 5::int \ :: A\ (Note that this way to construct an instance is not necessarily computable \<^enum> \<^term>\\tag_attribute = X, level = Y, x = Z\ :: A\ \<^enum> \<^term>\\tag_attribute = X, level = Y, x = Z, \ = M\ :: ('a A_scheme)\ \<^enum> there is an entire proof infra-structure allowing to reason about \<^emph>\doc\_class instances\; this involves the constructor, the selectors (representing the \<^emph>\attributes\ in OO lingo) the update functions, the rules to establish equality and, if possible the code generator setups: \<^enum> \<^term>\A.make :: int \ int option \ int \ A\ \<^enum> \<^term>\A.level :: 'a A_scheme \ int option\ \<^enum> \<^term>\A.level_update :: (int option \ int option) \ 'a A_scheme \ 'a A_scheme\ \<^enum> ... together with the rules such as: \<^enum> @{thm [display] A.simps(2)} \<^enum> @{thm [display] A.simps(6)} \<^enum> ... \ text\The generated theory of the \<^theory_text>\doc_class\ A can be inspected, of course, by:\ find_theorems (60) name:Conceptual name:A text\A more abstract view on the state of the DOF machine can be found here:\ print_doc_classes print_doc_items text\... and an ML-level output:\ ML\ val docitem_tab = DOF_core.get_instances \<^context>; val isa_transformer_tab = DOF_core.get_isa_transformers \<^context>; val docclass_tab = DOF_core.get_onto_classes \<^context>; \ ML\ Name_Space.dest_table docitem_tab; Name_Space.dest_table isa_transformer_tab; Name_Space.dest_table docclass_tab; \ text\... or as ML assertion: \ ML\ @{assert} (Name_Space.dest_table docitem_tab = []); fun match ("Conceptual.A", (* the long-name *) DOF_core.Onto_Class {params, name, virtual,inherits_from=NONE, attribute_decl, rejectS=[],rex=[], invs=[]}) = (Binding.name_of name = "A") | match _ = false; @{assert} (exists match (Name_Space.dest_table docclass_tab)) \ text\As a consequence of the theory of the \<^theory_text>\doc_class\ \A\, the code-generator setup lets us evaluate statements such as: \ value\ the(A.level (A.make 3 (Some 4) 5)) = 4\ text\And further, as a consequence of the above semantic construction of \<^theory_text>\doc_class\'es, the internal \\\-calculus representation of class instances looks as follows:\ ML\ @{term \the(A.level (A.make 3 (Some 4) 5))\}; fun match (Const("Option.option.the",_) $ (Const ("Conceptual.A.level",_) $ (Const ("Conceptual.A.make", _) $ u $ v $ w))) = true |match _ = false; @{assert} (match @{term \the(A.level (A.make 3 (Some 4) 5))\}) \ text\And finally, via the code-generation, we have the following programmable access to values representing class instances:\ ML\ val A_make = @{code A.make}; val zero = @{code "0::int"}; val one = @{code "1::int"}; val add = @{code "(+) :: int \ int \ int"}; A_make zero (SOME one) (add one one) \ section\Building up a conceptual class hierarchy:\ text\An independent class-tree root: \ doc_class B = level :: "int option" x :: "string" (* attributes live in their own name-space *) y :: "string list" <= "[]" (* and can have arbitrary type constructors *) (* LaTeX may have problems with this, though *) text\We may even use type-synonyms for class synonyms ...\ type_synonym XX = B section\Examples of inheritance \ doc_class C = XX + z :: "A option" <= None (* A LINK, i.e. an attribute that has a type referring to a document class. Mathematical relations over document items can be modeled. *) g :: "thm" (* a reference to the proxy-type 'thm' allowing to denote references to theorems inside attributes *) datatype enum = X1 | X2 | X3 (* we add an enumeration type ... *) doc_class D = B + x :: "string" <= "\def \\" (* overriding default *) a1 :: enum <= "X2" (* class - definitions may be mixed with arbitrary HOL-commands, thus also local definitions of enumerations *) a2 :: int <= 0 doc_class E = D + x :: "string" <= "''qed''" (* overriding default *) doc_class F = properties :: "term list" r :: "thm list" u :: "file" s :: "typ list" b :: "(A \ C) set" <= "{}" (* This is a relation link, roughly corresponding to an association class. It can be used to track claims to result - relations, for example.*) b' :: "(A \ C) list" <= "[]" invariant br :: "r \ \ [] \ card(b \) \ 3" and br':: "r \ \ [] \ length(b' \) \ 3" and cr :: "properties \ \ []" text\The effect of the invariant declaration is to provide intern HOL definitions for validation functions of this invariant. They can be referenced as follows:\ thm br_inv_def thm br'_inv_def thm cr_inv_def term "\F.tag_attribute = 5, properties = [], r = [], u = undefined, s = [], b = {}, b' = []\" term "br'_inv (\F.tag_attribute = 5, properties = [], r = [], u = undefined, s = [], b = {}, b' = []\) " text\Now, we can use these definitions in order to generate code for these validation functions. Note, however, that not everything that we can write in an invariant (basically: HOL) is executable, or even compilable by the code generator setup:\ ML\ val cr_inv_code = @{code "cr_inv"} \ \ \works albeit thm is abstract ...\ text\while in :\ ML\ val br_inv_code = @{code "br_inv"} \ \\this does not work ...\ text\... the compilation fails due to the fact that nothing prevents the user to define an infinite relation between \<^typ>\A\ and \<^typ>\C\. However, the alternative variant: \ ML\ val br'_inv_code = @{code "br'_inv"} \ \ \does work ...\ text\... is compilable ...\ doc_class G = C + g :: "thm" <= "@{thm \HOL.refl\}" (* warning overriding attribute expected*) doc_class M = ok :: "unit" accepts "A ~~ \C || D\\<^sup>* ~~ \F\" text\The final class and item tables look like this:\ print_doc_classes print_doc_items ML\ map fst (Name_Space.dest_table (DOF_core.get_onto_classes \<^context>)); let val class_ids_so_far = ["Conceptual.A", "Conceptual.B", "Conceptual.C", "Conceptual.D", "Conceptual.E", "Conceptual.F", "Conceptual.G", "Conceptual.M", "Isa_COL.figure", "Isa_COL.chapter", "Isa_COL.figure2", "Isa_COL.section", "Isa_COL.paragraph", "Isa_COL.subsection", "Isa_COL.figure_group", "Isa_COL.text_element", "Isa_COL.subsubsection", "Isa_COL.side_by_side_figure"] val docclass_tab = map fst (Name_Space.dest_table (DOF_core.get_onto_classes \<^context>)); in @{assert} (class_ids_so_far = docclass_tab) end\ section\For Test and Validation\ text*[sdf] \ Lorem ipsum ... \ \ \anonymous reference\ text*[sdfg :: F] \ Lorem ipsum ...\ \ \some F instance \ end