theory BAC2017 imports "../../../ontologies/mathex_onto" Deriv Transcendental begin open_monitor*[exam::MathExam] (* currently rethinking on "deep ontologies" necessary ... Achim text*[idir::Author,affiliation="''LRI, CentraleSupelec''", email="''idir.aitsadoune@centralesupelec.fr''"] {*Idir AIT SADOUNE*} text*[keller::Author,affiliation="''LRI, Univ. Paris-Sud''", email="''Chantal.Keller@lri.fr''"] {*Chantal KELLER*} text{* This example is an excerpt from the french baccaleareat 2017. The textual explanations were kept in french. *} *) text*[header::Header,examSubject="[analysis,geometry]", date="''21-06-2017''", timeAllowed="240::int"]{* BACCALAUREAT GENERAL MATHEMATIQUES *} text{* \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, qu’il 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 l’appréciation des copies. \end{itemize} *} text*[exo1 :: Exercise, concerns= "{setter,student,checker,externalExaminer}"] {* On considère la fonction h définie sur l’intervalle [0..+\] par : @{term "h(x) = x * exponent (-x)"} *} definition h :: "real \ real" where "h x \ x * exp (- x)" text*[q1::Task, concerns= "{setter,student}", level="oneStar", mark="1::int", type="formal"] {* Déterminer la limite de la fonction @{term h} en +\. *} text*[a1::Answer_Formal_Step] {* Fill in term and justification*} lemma q1 : "(h \ (0::real)) at_top" sorry text*[v1::Validation, proofs="[@{thm ''HOL.refl''}::thm]"] {* See lemma @{thm q1}. *} text*[q2::Task, concerns= "{setter,checker,student}", level="oneStar", mark="1::int", type="formal"] {* Étudier les variations de la fonction @{term h} sur l'intervalle [0..+\] et dresser son tableau de variation *} text*[a2::Answer_Formal_Step] {* Fill in term and justification*} definition h' :: "real \ real" where "h' x \ (1 - x) * exp (- x)" lemma q2_a : "DERIV h x :> h' x" proof - have * : "DERIV (exp \ uminus) x :> - (exp (-x))" sorry (* by (simp add: has_derivative_compose) *) have ** : "DERIV id x :> 1" by (metis DERIV_ident eq_id_iff) have *** : "DERIV h x :> x * (- (exp (- x))) + 1 * (exp (- x))" sorry (* by (simp add: * ** has_derivative_mult comp_def) *) show ?thesis sorry (* by (metis "***" left_diff_distrib mult_minus_right uminus_add_conv_diff) *) qed lemma q2_b : "0 \ x \ x \ y \ y \ 1 \ h x \ h y" sorry lemma q2_c : "1 \ x \ x \ y \ h x \ h y" sorry text*[v2::Validation, proofs="[@{thm ''BAC2017.q2_b''}, @{thm ''BAC2017.q2_c''}]"] {* See lemmas @{thm q2_b} and @{thm q2_c}. *} text*[q3a::Task, concerns= "{setter,checker,student}", level="oneStar", mark="1::int", type="formal"] {* Vérifier que pour tout nombre réel x appartenant à l'intervalle [0..+\], on a : @{term "h x = (exp (- x)) - (h' x)"}. *} text*[a3a::Answer_Formal_Step] {* Fill in term and justification*} lemma q3a : "h x = (exp (- x)) - (h' x)" by (simp add : h_def h'_def left_diff_distrib) subsubsection*[v3a::Validation, proofs="[@{thm ''BAC2017.q3a''}::thm]"] {* See lemma @{thm q3a}. *} subsection*[sol1 :: Solution, content="[exo1::Exercise]", valids = "[v1::Validation,v2,v3a]"] {* See validations. *} close_monitor*[exam] end