theory InnerSyntaxAntiquotations imports "../../ontologies/Conceptual" begin ML\ val ({tab = x, ...},y,z)= DOF_core.get_data @{context}; Symtab.dest z; \ lemma murks : "T = {ert,dfg}" sorry text\Example for a meta-attribute of ODL-type @{typ "file"} with an appropriate ISA for the file @{file "./Attributes.thy"}\ text*[xcv::F, u="@{file ''./examples/conceptual/Attributes.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 \verb+F+ with a relational link, and links to formal Isabelle items. \ text*[xcv4::F, r="[@{thm ''HOL.refl''}, @{thm ''InnerSyntaxAntiquotations.murks''}]", b="{(@{docitem ''xcv1''},@{docitem ''xcv2''})}", s="[@{typ ''int list''}]"]\Lorem ipsum ...\ text\... and here we add a relation between @{docitem \xcv3\} and @{docitem \xcv2\} into the relation \verb+b+ of @{docitem_ref \xcv4\} \ update_instance*[xcv4::F, b+="{(@{docitem ''xcv3''},@{docitem ''xcv2''})}"] text\And here is the result on term level:\ ML\ @{docitem_attr b::xcv4} \ end