Isabelle_DOF/examples/simple/MathExam.thy

35 lines
976 B
Plaintext
Raw Normal View History

2018-04-17 14:43:08 +00:00
theory MathExam
imports "../../ontologies/mathex_onto"
begin
2018-04-17 14:43:08 +00:00
open_monitor*[exam::MathExam]
text*[idir::Author, affiliation="''CentraleSupelec''",
email="''idir.aitsadoune@centralesupelec.fr''"]
2018-04-27 09:33:47 +00:00
{*Idir AIT SADOUNE*}
2018-04-17 14:43:08 +00:00
text*[header::Header,examSubject= "algebra", examTitle="''Exam number 1''"]
{* Please follow directions carefully and show all your work.*}
2018-04-18 07:36:21 +00:00
section*[exo1 :: Exercise, Exercise.content="[q1::Question,q2,q3]", mark="15::int"]{* Exercise 1.*}
text*[q1::Question, level="twoStars", Question.mark="5::int"]
2018-04-17 14:43:08 +00:00
{*
Give an example of each of the following :
\begin{itemize}
\item - a rational number which is not integer.
\item - a real number which is not rational.
\end{itemize}
2018-04-17 14:43:08 +00:00
*}
text*[q2::Question, level="oneStar", Question.mark="5::int"]
2018-04-17 14:43:08 +00:00
{*
2018-04-27 09:15:27 +00:00
Write in interval notation : @{term "-3 < x"} and @{term "x < 5"}
2018-04-17 14:43:08 +00:00
*}
text*[q3::Question, level="oneStar", Question.mark="5::int"]
2018-04-27 09:15:27 +00:00
{* True or false : @{term "0/8 = 0"} *}
2018-04-17 14:43:08 +00:00
2018-04-18 07:36:21 +00:00
close_monitor*[exam]
end