chapter\Setting and modifying attributes of doc-items\ theory Concept_Example imports "../ontologies/Conceptual/Conceptual" (* we use the generic "Conceptual" ontology *) begin text\@{theory "Isabelle_DOF-tests.Conceptual"} provides a monitor @{typ M} enforcing a particular document structure. Here, we say: From now on, this structural rules are respected wrt. all doc\_classes M is enabled for.\ open_monitor*[struct::M] section*[a::A, x = "3"] \ Lorem ipsum dolor sit amet, ... \ text*[c1::C, x = "''beta''"] \ ... suspendisse non arcu malesuada mollis, nibh morbi, ... \ text*[d::D, a1 = "X3"] \ ... phasellus amet id massa nunc, pede suscipit repellendus, ... @{docitem "c1"} @{thm "refl"} \ update_instance*[d::D, a1 := X2] text\ ... in ut tortor ... @{docitem \a\} ... @{A \a\}\ text*[c2::C, x = "\delta\"] \ ... in ut tortor eleifend augue pretium consectetuer. \ text\Note that both the notations @{term "''beta''"} and @{term "\beta\"} are possible; the former is a more ancient format only supporting pure ascii, while the latter also supports fancy unicode such as: @{term "\\\<^sub>i''\"} \ text*[f::F] \ Lectus accumsan velit ultrices, ... }\ theorem some_proof : "True" by simp text\This is an example where we add a theorem into a kind of "result-list" of the doc-item f.\ update_instance*[f::F,r:="[@{thm ''Concept_Example.some_proof''}]"] text\ ..., mauris amet, id elit aliquam aptent id, ... @{docitem \a\} \ text\Here we add and maintain a link that is actually modeled as m-to-n relation ...\ update_instance*[f::F,b:="{(@{docitem \a\}::A,@{docitem \c1\}::C), (@{docitem \a\}, @{docitem \c2\})}"] close_monitor*[struct] text\And the trace of the monitor is:\ ML\@{trace_attribute struct}\ print_doc_classes print_doc_items end