Isabelle_DOF/examples/math_exam/BAC2017.thy

72 lines
2.5 KiB
Plaintext
Raw Normal View History

2018-06-04 11:44:08 +00:00
theory BAC2017
imports "../../ontologies/mathex_onto"
2018-06-05 18:56:02 +00:00
Deriv Transcendental
2018-06-04 11:44:08 +00:00
begin
open_monitor*[exam::MathExam]
section*[idir::Author,affiliation="''CentraleSupelec''",
email="''idir.aitsadoune@centralesupelec.fr''"]
{*Idir AIT SADOUNE*}
2018-06-05 17:20:00 +00:00
subsection*[header::Header,examSubject= "[algebra,geometry]",
2018-06-04 11:44:08 +00:00
examTitle="''BACCALAUREAT GENERAL MATHEMATIQUES''",
date="''21-06-2017''",
timeAllowed="240::int"]
{*
\begin{itemize}
\item Les calculatrices électroniques de poche sont autorisées, conformément à la réglementation en vigueur.
\item Le sujet est composé de 4 exercices indépendants.
\item Le candidat doit traiter tous les exercices.
\item Le candidat est invité à faire figurer sur la copie toute trace de recherche, même incomplète ou non fructueuse, quil aura développée.
\item Il est rappelé que la qualité de la rédaction, la clarté et la précision des raisonnements entreront pour une part importante dans lappréciation des copies.
\end{itemize}
*}
2018-06-05 17:20:00 +00:00
subsubsection*[exo1 :: Exercise,Exercise.concerns= "{examiner,validator,student}",
2018-06-04 11:44:08 +00:00
Exercise.content="[q1::Task]"]
{*
2018-06-05 17:20:00 +00:00
On considère la fonction h définie sur lintervalle [0..+\<infinity>] par : @{term "h(x) = x * exponent (-x)"}
2018-06-04 11:44:08 +00:00
*}
2018-06-05 18:56:02 +00:00
definition h :: "real \<Rightarrow> real"
where "h x \<equiv> x * exp (- x)"
2018-06-04 11:44:08 +00:00
2018-06-05 17:20:00 +00:00
subsubsection*[q1::Task, Task.concerns= "{examiner,validator,student}",
2018-06-04 11:44:08 +00:00
level="oneStar", mark="1::int", type="formal"]
2018-06-05 18:56:02 +00:00
{* Déterminer la limite de la fonction @{term h} en +\<infinity>. *}
2018-06-04 11:44:08 +00:00
2018-06-05 17:20:00 +00:00
text*[a1::Answer_Formal_Step]
{* First Step: Fill in term and justification*}
2018-06-04 11:44:08 +00:00
2018-06-05 18:56:02 +00:00
lemma q1 : "(h \<longlongrightarrow> (0::real)) at_top"
sorry
subsubsection*[q2::Task, Task.concerns= "{examiner,validator,student}",
level="oneStar", mark="1::int", type="formal"]
{* Étudier les variations de la fonction @{term h} sur l'intervalle [0..+\<infinity>] et dresser son tableau
de variation} *}
find_theorems exp
text*[a2::Answer_Formal_Step]
{* First Step: Fill in term and justification*}
lemma q2_a : "DERIV h x :> (1 - x) * exp (- x)"
proof -
have * : "DERIV (exp \<circ> uminus) x :> - (exp (-x))"
by (simp add: has_derivative_minus has_derivative_compose Transcendental.DERIV_exp)
have ** : "DERIV id x :> 1"
by (metis DERIV_ident eq_id_iff)
have *** : "DERIV h x :> x * (- (exp (- x))) + 1 * (exp (- x))"
by (simp add: * ** has_derivative_mult)
show ?thesis
by (metis "***" left_diff_distrib mult_minus_right uminus_add_conv_diff)
qed
2018-06-04 11:44:08 +00:00
close_monitor*[exam]
end