theory Example imports "../../ontologies/CENELEC_50126" begin section{* Some show-off's of general antiquotations. *} (* some show-off of standard anti-quotations: *) text \THIS IS A TEXT\ text\ @{thm refl} of name @{thm [source] refl} @{thm[mode=Rule] conjI} @{file "../../Isa_DOF.thy"} @{value "3+4::int"} @{const hd} @{theory List}} @{term "3"} @{type bool} @{term [show_types] "f x = a + x"} \ section{* Example *} text*[tralala] {* Brexit means Brexit *} text*[ass1::assumption] {* Brexit means Brexit *} text*[hyp1::hypothesis] {* P inequal NP *} text*[ass122::srac] {* The overall sampling frequence of the odometer subsystem is therefore 14 khz, which includes sampling, computing and result communication times... *} text*[t10::test_result] {* This is a meta-test. This could be an ML-command that governs the external test-execution via, eg., a makefile or specific calls to a test-environment or test-engine *} text \ As established by @{docref (unchecked) \t10\}, @{docref (define) \t10\} \ text \ the @{docref \t10\} as well as the @{docref \ass122\}\ text \ represent a justification of the safety related applicability condition @{srac \ass122\} aka exported constraint @{ec \ass122\}. \ section{* Some Tests for Ontology Framework and its CENELEC Instance *} declare_reference* [lalala::requirement, alpha="main", beta=42] declare_reference* [blablabla::cid, alpha="beta sdf", beta=gamma, delta=dfg_fgh\<^sub>1] print_doc_classes print_doc_items paragraph* [sdfk] \ just a paragraph - lexical variant \ section{* Text Antiquotation Infrastructure ... *} text{* @{docref \lalala\} -- produces warning. *} text{* @{docref (unchecked) \lalala\} -- produces no warning. *} text{* @{docref \ass122\} -- global reference to a text-item in another file. *} text{* @{ec \ass122\} -- global reference to a exported constraint in another file. Note that the link is actually a srac, which, according to the ontology, happens to be an "ec". *} end