Isabelle_DOF/ontologies/mathex_onto.thy

37 lines
612 B
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"
2018-04-27 09:33:47 +00:00
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
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 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