Isabelle_DOF/ontologies/mathex_onto.thy

55 lines
1.3 KiB
Plaintext
Raw Normal View History

2018-04-06 12:25:21 +00:00
theory mathex_onto
imports "../Isa_DOF"
begin
2018-04-17 14:42:45 +00:00
doc_class Author =
affiliation :: "string"
email :: "string"
2018-04-06 12:32:22 +00:00
2018-04-17 14:42:45 +00:00
datatype Subject =
algebra | geometry | statistical
2018-04-06 12:32:22 +00:00
2018-04-17 14:42:45 +00:00
datatype Level =
oneStar | twoStars | threeStars
2018-04-06 12:32:22 +00:00
text{* In our scenario, content has three different types of addressees:
\<^item> the @{emph \<open>examiner\<close>}, i.e. the author of the exam,
\<^item> the @{emph \<open>student\<close>}, i.e. the addressee of the exam,
\<^item> the @{emph \<open>validator\<close>}, i.e. a person that checks the exam for
feasibility and non-ambiguity.
Note that the latter quality assurance mechanism is used in many universities,
where for organizational reasons the execution of an exam takes place in facilities
where the author of the exam is not expected to be physically present.
*}
datatype ContentClass =
examiner | validator | student
2018-04-17 14:42:45 +00:00
datatype Grade =
A1 | A2 | A3
2018-04-06 12:32:22 +00:00
2018-04-17 14:42:45 +00:00
doc_class Header =
2018-04-18 07:36:21 +00:00
examTitle :: string
2018-04-17 14:42:45 +00:00
examGrade :: Grade
examSubject :: Subject
2018-04-17 09:14:52 +00:00
2018-04-17 14:42:45 +00:00
doc_class Question =
level :: Level
mark :: integer
2018-04-17 14:42:45 +00:00
doc_class Exercise=
content :: "(Question) list"
2018-04-27 09:15:27 +00:00
mark :: integer
2018-04-17 14:42:45 +00:00
doc_class Solution =
content :: "(Question) list"
2018-04-17 14:42:45 +00:00
doc_class MathExam=
content :: "(Header + Author + Exercise) list"
2018-04-18 07:36:21 +00:00
where "((Author)+ ~
Header ~
(Exercise)+ )"
2018-04-06 12:25:21 +00:00
end