src/Tools/isac/Knowledge/Diff.thy
author wenzelm
Fri, 11 Jun 2021 11:49:34 +0200
changeset 60294 6623f5cdcb19
parent 60291 52921aa0e14a
child 60296 81b6519da42b
permissions -rw-r--r--
ML antiquotation for formally checked Rule.Eval;
     1 (* differentiation over the reals
     2    author: Walther Neuper
     3    000516   
     4  *)
     5 
     6 theory Diff imports Calculus Trig LogExp Rational Root Poly Base_Tools begin
     7 
     8 ML \<open>
     9 @{term "sin x"}
    10 \<close>
    11 
    12 consts
    13 
    14   d_d           :: "[real, real]=> real"
    15 
    16   (*descriptions in the related problems*)
    17   derivativeEq  :: "bool => una"
    18 
    19   (*predicates*)
    20   primed        :: "'a => 'a" (*"primed A" -> "A'"*)
    21 
    22   (*the CAS-commands, eg. "Diff (2*x \<up> 3, x)", 
    23 			  "Differentiate (A = s * (a - s), s)"*)
    24   Diff           :: "[real * real] => real"
    25   Differentiate  :: "[bool * real] => bool"
    26 
    27   (*subproblem-name*)
    28   differentiate  :: "[char list * char list list * char list, real, real] => real"
    29                	   ("(differentiate (_)/ (_ _ ))" 9)
    30 
    31 text \<open>a variant of the derivatives defintion:
    32 
    33   d_d            :: "(real => real) => (real => real)"
    34 
    35   advantages:
    36 (1) no variable 'bdv' on the meta-level required
    37 (2) chain_rule "d_d (%x. (u (v x))) = (%x. (d_d u)) (v x) * d_d v"
    38 (3) and no specialized chain-rules required like
    39     diff_sin_chain "d_d bdv (sin u)    = cos u * d_d bdv u"
    40 
    41   disadvantage: d_d (%x. 1 + x^2) = ... differs from high-school notation
    42 \<close>
    43 
    44 axiomatization where (*stated as axioms, todo: prove as theorems
    45         'bdv' is a constant on the meta-level  *)
    46   diff_const:     "[| Not (bdv occurs_in a) |] ==> d_d bdv a = 0" and
    47   diff_var:       "d_d bdv bdv = 1" and
    48   diff_prod_const:"[| Not (bdv occurs_in u) |] ==>  
    49 					 d_d bdv (u * v) = u * d_d bdv v" and
    50 
    51   diff_sum:       "d_d bdv (u + v)     = d_d bdv u + d_d bdv v" and
    52   diff_dif:       "d_d bdv (u - v)     = d_d bdv u - d_d bdv v" and
    53   diff_prod:      "d_d bdv (u * v)     = d_d bdv u * v + u * d_d bdv v" and
    54   diff_quot:      "Not (v = 0) ==> (d_d bdv (u / v) =  
    55 	           (d_d bdv u * v - u * d_d bdv v) / v \<up> 2)" and
    56 
    57   diff_sin:       "d_d bdv (sin bdv)   = cos bdv" and
    58   diff_sin_chain: "d_d bdv (sin u)     = cos u * d_d bdv u" and
    59   diff_cos:       "d_d bdv (cos bdv)   = - sin bdv" and
    60   diff_cos_chain: "d_d bdv (cos u)     = - sin u * d_d bdv u" and
    61   diff_pow:       "d_d bdv (bdv \<up> n) = n * (bdv \<up> (n - 1))" and
    62   diff_pow_chain: "d_d bdv (u \<up> n)   = n * (u \<up> (n - 1)) * d_d bdv u" and
    63   diff_ln:        "d_d bdv (ln bdv)    = 1 / bdv" and
    64   diff_ln_chain:  "d_d bdv (ln u)      = d_d bdv u / u" and
    65   diff_exp:       "d_d bdv (exp bdv)   = exp bdv" and
    66   diff_exp_chain: "d_d bdv (exp u)     = exp u * d_d x u" and
    67 (*
    68   diff_sqrt      "d_d bdv (sqrt bdv)  = 1 / (2 * sqrt bdv)"
    69   diff_sqrt_chain"d_d bdv (sqrt u)    = d_d bdv u / (2 * sqrt u)"
    70 *)
    71   (*...*)
    72 
    73   frac_conv:       "[| bdv occurs_in b; 0 < n |] ==>  
    74 		    a / (b \<up> n) = a * b \<up> (-n)" and
    75   frac_sym_conv:   "n < 0 ==> a * b \<up> n = a / b \<up> (-n)" and
    76 
    77   sqrt_conv_bdv:   "sqrt bdv = bdv \<up> (1 / 2)" and
    78   sqrt_conv_bdv_n: "sqrt (bdv \<up> n) = bdv \<up> (n / 2)" and
    79 (*Ambiguous input\<^here> produces 3 parse trees -----------------------------\\*)
    80   sqrt_conv:       "bdv occurs_in u ==> sqrt u = u \<up> (1 / 2)" and
    81 (*Ambiguous input\<^here> produces 3 parse trees -----------------------------//*)
    82   sqrt_sym_conv:   "u \<up> (a / 2) = sqrt (u \<up> a)" and
    83 
    84   root_conv:       "bdv occurs_in u ==> nroot n u = u \<up> (1 / n)" and
    85   root_sym_conv:   "u \<up> (a / b) = nroot b (u \<up> a)" and
    86 
    87   realpow_pow_bdv: "(bdv \<up> b) \<up> c = bdv \<up> (b * c)"
    88 
    89 ML \<open>
    90 (** eval functions **)
    91 
    92 fun primed (Const (id, T)) = Const (id ^ "'", T)
    93   | primed (Free (id, T)) = Free (id ^ "'", T)
    94   | primed t = raise ERROR ("primed called with arg = '"^ UnparseC.term t ^"'");
    95 
    96 (*("primed", ("Diff.primed", eval_primed "#primed"))*)
    97 fun eval_primed _ _ (p as (Const ("Diff.primed",_) $ t)) _ =
    98     SOME ((UnparseC.term p) ^ " = " ^ UnparseC.term (primed t),
    99 	  HOLogic.Trueprop $ (TermC.mk_equality (p, primed t)))
   100   | eval_primed _ _ _ _ = NONE;
   101 \<close>
   102 setup \<open>KEStore_Elems.add_calcs
   103   [("primed", ("Diff.primed", eval_primed "#primed"))]\<close>
   104 ML \<open>
   105 (** rulesets **)
   106 
   107 (*.converts a term such that differentiation works optimally.*)
   108 val diff_conv =   
   109     Rule_Def.Repeat {id="diff_conv", 
   110 	 preconds = [], 
   111 	 rew_ord = ("termlessI",termlessI), 
   112 	 erls = Rule_Set.append_rules "erls_diff_conv" Rule_Set.empty 
   113 			   [\<^rule_eval>\<open>Prog_Expr.occurs_in\<close> (Prog_Expr.eval_occurs_in ""),
   114 			    Rule.Thm ("not_true",ThmC.numerals_to_Free @{thm not_true}),
   115 			    Rule.Thm ("not_false",ThmC.numerals_to_Free @{thm not_false}),
   116 			    \<^rule_eval>\<open>less\<close> (Prog_Expr.eval_equ "#less_"),
   117 			    Rule.Thm ("and_true",ThmC.numerals_to_Free @{thm and_true}),
   118 			    Rule.Thm ("and_false",ThmC.numerals_to_Free @{thm and_false})
   119 			    ], 
   120 	 srls = Rule_Set.Empty, calc = [], errpatts = [],
   121 	 rules =
   122   [Rule.Thm ("frac_conv", ThmC.numerals_to_Free @{thm frac_conv}),
   123      (*"?bdv occurs_in ?b \<Longrightarrow> 0 < ?n \<Longrightarrow> ?a / ?b \<up> ?n = ?a * ?b \<up> - ?n"*)
   124 		   Rule.Thm ("sqrt_conv_bdv", ThmC.numerals_to_Free @{thm sqrt_conv_bdv}),
   125 		     (*"sqrt ?bdv = ?bdv \<up> (1 / 2)"*)
   126 		   Rule.Thm ("sqrt_conv_bdv_n", ThmC.numerals_to_Free @{thm sqrt_conv_bdv_n}),
   127 		     (*"sqrt (?bdv \<up> ?n) = ?bdv \<up> (?n / 2)"*)
   128 		   Rule.Thm ("sqrt_conv", ThmC.numerals_to_Free @{thm sqrt_conv}),
   129 		     (*"?bdv occurs_in ?u \<Longrightarrow> sqrt ?u = ?u \<up> (1 / 2)"*)
   130 		   Rule.Thm ("root_conv", ThmC.numerals_to_Free @{thm root_conv}),
   131 		     (*"?bdv occurs_in ?u \<Longrightarrow> nroot ?n ?u = ?u \<up> (1 / ?n)"*)
   132 		   Rule.Thm ("realpow_pow_bdv", ThmC.numerals_to_Free @{thm realpow_pow_bdv}),
   133 		     (* "(?bdv \<up> ?b) \<up> ?c = ?bdv \<up> (?b * ?c)"*)
   134 		   \<^rule_eval>\<open>times\<close> (**)(eval_binop "#mult_"),
   135 		   Rule.Thm ("rat_mult",ThmC.numerals_to_Free @{thm rat_mult}),
   136 		     (*a / b * (c / d) = a * c / (b * d)*)
   137 		   Rule.Thm ("times_divide_eq_right",ThmC.numerals_to_Free @{thm times_divide_eq_right}),
   138 		     (*?x * (?y / ?z) = ?x * ?y / ?z*)
   139 		   Rule.Thm ("times_divide_eq_left",ThmC.numerals_to_Free @{thm times_divide_eq_left})
   140 		     (*?y / ?z * ?x = ?y * ?x / ?z*)
   141 		 ],
   142 	 scr = Rule.Empty_Prog};
   143 \<close>
   144 ML \<open>
   145 (*.beautifies a term after differentiation.*)
   146 val diff_sym_conv =   
   147     Rule_Def.Repeat {id="diff_sym_conv", 
   148 	 preconds = [], 
   149 	 rew_ord = ("termlessI",termlessI), 
   150 	 erls = Rule_Set.append_rules "erls_diff_sym_conv" Rule_Set.empty 
   151 			   [\<^rule_eval>\<open>less\<close> (Prog_Expr.eval_equ "#less_")], 
   152 	 srls = Rule_Set.Empty, calc = [], errpatts = [],
   153 	 rules = [Rule.Thm ("frac_sym_conv", ThmC.numerals_to_Free @{thm frac_sym_conv}),
   154 		  Rule.Thm ("sqrt_sym_conv", ThmC.numerals_to_Free @{thm sqrt_sym_conv}),
   155 		  Rule.Thm ("root_sym_conv", ThmC.numerals_to_Free @{thm root_sym_conv}),
   156 		  Rule.Thm ("sym_real_mult_minus1",
   157 		       ThmC.numerals_to_Free (@{thm real_mult_minus1} RS @{thm sym})),
   158 		      (*- ?z = "-1 * ?z"*)
   159 		  Rule.Thm ("rat_mult",ThmC.numerals_to_Free @{thm rat_mult}),
   160 		  (*a / b * (c / d) = a * c / (b * d)*)
   161 		  Rule.Thm ("times_divide_eq_right",ThmC.numerals_to_Free @{thm times_divide_eq_right}),
   162 		  (*?x * (?y / ?z) = ?x * ?y / ?z*)
   163 		  Rule.Thm ("times_divide_eq_left",ThmC.numerals_to_Free @{thm times_divide_eq_left}),
   164 		  (*?y / ?z * ?x = ?y * ?x / ?z*)
   165 		  \<^rule_eval>\<open>times\<close> (**)(eval_binop "#mult_")
   166 		 ],
   167 	 scr = Rule.Empty_Prog};
   168 
   169 (*..*)
   170 val srls_diff = 
   171     Rule_Def.Repeat {id="srls_differentiate..", 
   172 	 preconds = [], 
   173 	 rew_ord = ("termlessI",termlessI), 
   174 	 erls = Rule_Set.empty, 
   175 	 srls = Rule_Set.Empty, calc = [], errpatts = [],
   176 	 rules = [\<^rule_eval>\<open>Prog_Expr.lhs\<close> (Prog_Expr.eval_lhs "eval_lhs_"),
   177 		  \<^rule_eval>\<open>Prog_Expr.rhs\<close> (Prog_Expr.eval_rhs "eval_rhs_"),
   178 		  \<^rule_eval>\<open>Diff.primed\<close> (eval_primed "Diff.primed")
   179 		  ],
   180 	 scr = Rule.Empty_Prog};
   181 \<close>
   182 ML \<open>
   183 (*..*)
   184 val erls_diff = 
   185     Rule_Set.append_rules "erls_differentiate.." Rule_Set.empty
   186                [Rule.Thm ("not_true",ThmC.numerals_to_Free @{thm not_true}),
   187 		Rule.Thm ("not_false",ThmC.numerals_to_Free @{thm not_false}),
   188 		
   189 		\<^rule_eval>\<open>Prog_Expr.ident\<close> (Prog_Expr.eval_ident "#ident_"),
   190 		\<^rule_eval>\<open>Prog_Expr.is_atom\<close> (Prog_Expr.eval_is_atom "#is_atom_"),
   191 		\<^rule_eval>\<open>Prog_Expr.occurs_in\<close> (Prog_Expr.eval_occurs_in ""),
   192 		\<^rule_eval>\<open>Prog_Expr.is_const\<close> (Prog_Expr.eval_const "#is_const_")
   193 		];
   194 
   195 (*.rules for differentiation, _no_ simplification.*)
   196 val diff_rules =
   197     Rule_Def.Repeat {id="diff_rules", preconds = [], rew_ord = ("termlessI",termlessI), 
   198 	 erls = erls_diff, srls = Rule_Set.Empty, calc = [], errpatts = [],
   199 	 rules = [Rule.Thm ("diff_sum",ThmC.numerals_to_Free @{thm diff_sum}),
   200 		  Rule.Thm ("diff_dif",ThmC.numerals_to_Free @{thm diff_dif}),
   201 		  Rule.Thm ("diff_prod_const",ThmC.numerals_to_Free @{thm diff_prod_const}),
   202 		  Rule.Thm ("diff_prod",ThmC.numerals_to_Free @{thm diff_prod}),
   203 		  Rule.Thm ("diff_quot",ThmC.numerals_to_Free @{thm diff_quot}),
   204 		  Rule.Thm ("diff_sin",ThmC.numerals_to_Free @{thm diff_sin}),
   205 		  Rule.Thm ("diff_sin_chain",ThmC.numerals_to_Free @{thm diff_sin_chain}),
   206 		  Rule.Thm ("diff_cos",ThmC.numerals_to_Free @{thm diff_cos}),
   207 		  Rule.Thm ("diff_cos_chain",ThmC.numerals_to_Free @{thm diff_cos_chain}),
   208 		  Rule.Thm ("diff_pow",ThmC.numerals_to_Free @{thm diff_pow}),
   209 		  Rule.Thm ("diff_pow_chain",ThmC.numerals_to_Free @{thm diff_pow_chain}),
   210 		  Rule.Thm ("diff_ln",ThmC.numerals_to_Free @{thm diff_ln}),
   211 		  Rule.Thm ("diff_ln_chain",ThmC.numerals_to_Free @{thm diff_ln_chain}),
   212 		  Rule.Thm ("diff_exp",ThmC.numerals_to_Free @{thm diff_exp}),
   213 		  Rule.Thm ("diff_exp_chain",ThmC.numerals_to_Free @{thm diff_exp_chain}),
   214 (*
   215 		  Rule.Thm ("diff_sqrt",ThmC.numerals_to_Free @{thm diff_sqrt}),
   216 		  Rule.Thm ("diff_sqrt_chain",ThmC.numerals_to_Free @{thm diff_sqrt_chain}),
   217 *)
   218 		  Rule.Thm ("diff_const",ThmC.numerals_to_Free @{thm diff_const}),
   219 		  Rule.Thm ("diff_var",ThmC.numerals_to_Free @{thm diff_var})
   220 		  ],
   221 	 scr = Rule.Empty_Prog};
   222 \<close>
   223 ML \<open>
   224 (*.normalisation for checking user-input.*)
   225 val norm_diff = 
   226   Rule_Def.Repeat
   227     {id="norm_diff", preconds = [], rew_ord = ("termlessI",termlessI), 
   228      erls = Rule_Set.Empty, srls = Rule_Set.Empty, calc = [], errpatts = [],
   229      rules = [Rule.Rls_ diff_rules, Rule.Rls_ norm_Poly ],
   230      scr = Rule.Empty_Prog};
   231 \<close>
   232 rule_set_knowledge
   233   erls_diff = \<open>prep_rls' erls_diff\<close> and
   234   diff_rules = \<open>prep_rls' diff_rules\<close> and
   235   norm_diff = \<open>prep_rls' norm_diff\<close> and
   236   diff_conv = \<open>prep_rls' diff_conv\<close> and
   237   diff_sym_conv = \<open>prep_rls' diff_sym_conv\<close>
   238 
   239 (** problem types **)
   240 setup \<open>KEStore_Elems.add_pbts
   241   [(Problem.prep_input @{theory} "pbl_fun" [] Problem.id_empty (["function"], [], Rule_Set.empty, NONE, [])),
   242     (Problem.prep_input @{theory} "pbl_fun_deriv" [] Problem.id_empty
   243       (["derivative_of", "function"],
   244         [("#Given" ,["functionTerm f_f", "differentiateFor v_v"]),
   245           ("#Find"  ,["derivative f_f'"])],
   246         Rule_Set.append_rules "empty" Rule_Set.empty [],
   247         SOME "Diff (f_f, v_v)", [["diff", "differentiate_on_R"],
   248 			  ["diff", "after_simplification"]])),
   249     (*here "named" is used differently from Integration"*)
   250     (Problem.prep_input @{theory} "pbl_fun_deriv_nam" [] Problem.id_empty
   251       (["named", "derivative_of", "function"],
   252         [("#Given" ,["functionEq f_f", "differentiateFor v_v"]),
   253           ("#Find"  ,["derivativeEq f_f'"])],
   254         Rule_Set.append_rules "empty" Rule_Set.empty [],
   255         SOME "Differentiate (f_f, v_v)",
   256         [["diff", "differentiate_equality"]]))]\<close>
   257 
   258 ML \<open>
   259 (** CAS-commands **)
   260 
   261 (*.handle cas-input like "Diff (a * x^3 + b, x)".*)
   262 (* val (t, pairl) = strip_comb (str2term "Diff (a * x^3 + b, x)");
   263    val [Const ("Product_Type.Pair", _) $ t $ bdv] = pairl;
   264    *)
   265 fun argl2dtss [Const ("Product_Type.Pair", _) $ t $ bdv] =
   266     [((Thm.term_of o the o (TermC.parse \<^theory>)) "functionTerm", [t]),
   267      ((Thm.term_of o the o (TermC.parse \<^theory>)) "differentiateFor", [bdv]),
   268      ((Thm.term_of o the o (TermC.parse \<^theory>)) "derivative", 
   269       [(Thm.term_of o the o (TermC.parse \<^theory>)) "f_f'"])
   270      ]
   271   | argl2dtss _ = raise ERROR "Diff.ML: wrong argument for argl2dtss";
   272 \<close>
   273 setup \<open>KEStore_Elems.add_mets
   274     [MethodC.prep_input @{theory} "met_diff" [] MethodC.id_empty
   275       (["diff"], [],
   276         {rew_ord'="tless_true",rls'=Atools_erls,calc = [], srls = Rule_Set.empty, prls=Rule_Set.empty,
   277           crls = Atools_erls, errpats = [], nrls = norm_diff},
   278         @{thm refl})]
   279 \<close>
   280 
   281 partial_function (tailrec) differentiate_on_R :: "real \<Rightarrow> real \<Rightarrow> real"
   282   where
   283 "differentiate_on_R f_f v_v = (
   284   let
   285     f_f' = Take (d_d v_v f_f)
   286   in (
   287     (Try (Rewrite_Set_Inst [(''bdv'',v_v)] ''diff_conv'')) #> (
   288     Repeat (
   289       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_sum'')) Or
   290       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_prod_const'')) Or
   291       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_prod'')) Or
   292       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_quot'')) Or
   293       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_sin'')) Or
   294       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_sin_chain'')) Or
   295       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_cos'')) Or
   296       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_cos_chain'')) Or
   297       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_pow'')) Or
   298       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_pow_chain'')) Or
   299       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_ln'')) Or
   300       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_ln_chain'')) Or
   301       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_exp'')) Or
   302       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_exp_chain'')) Or
   303       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_const'')) Or
   304       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_var'')) Or
   305       (Repeat (Rewrite_Set ''make_polynomial'')))) #> (
   306     Try (Rewrite_Set_Inst [(''bdv'',v_v)] ''diff_sym_conv''))
   307     ) f_f')"
   308 setup \<open>KEStore_Elems.add_mets
   309     [MethodC.prep_input @{theory} "met_diff_onR" [] MethodC.id_empty
   310       (["diff", "differentiate_on_R"],
   311         [("#Given" ,["functionTerm f_f", "differentiateFor v_v"]),
   312           ("#Find"  ,["derivative f_f'"])],
   313         {rew_ord'="tless_true", rls' = erls_diff, calc = [], srls = Rule_Set.empty, prls=Rule_Set.empty,
   314           crls = Atools_erls, errpats = [], nrls = norm_diff},
   315         @{thm differentiate_on_R.simps})]
   316 \<close>
   317 
   318 partial_function (tailrec) differentiateX :: "real \<Rightarrow> real \<Rightarrow> real"
   319   where
   320 "differentiateX f_f v_v = (
   321   let
   322     f_f' = Take (d_d v_v f_f)
   323   in (
   324     Repeat (
   325       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_sum'')) Or
   326       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_prod_const'' )) Or
   327       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_prod'')) Or
   328       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_quot'')) Or
   329       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_sin'')) Or
   330       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_sin_chain'')) Or
   331       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_cos'')) Or
   332       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_cos_chain'')) Or
   333       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_pow'')) Or
   334       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_pow_chain'')) Or
   335       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_ln'')) Or
   336       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_ln_chain'')) Or
   337       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_exp'')) Or
   338       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_exp_chain'')) Or
   339       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_const'')) Or
   340       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_var'')) Or
   341       (Repeat (Rewrite_Set ''make_polynomial'')))
   342     ) f_f')"
   343 setup \<open>KEStore_Elems.add_mets
   344     [MethodC.prep_input @{theory} "met_diff_simpl" [] MethodC.id_empty
   345       (["diff", "diff_simpl"],
   346         [("#Given", ["functionTerm f_f", "differentiateFor v_v"]),
   347          ("#Find" , ["derivative f_f'"])],
   348         {rew_ord'="tless_true", rls' = erls_diff, calc = [], srls = Rule_Set.empty, prls=Rule_Set.empty,
   349           crls = Atools_erls, errpats = [], nrls = norm_diff},
   350         @{thm differentiateX.simps})]
   351 \<close>
   352 
   353 partial_function (tailrec) differentiate_equality :: "bool \<Rightarrow> real \<Rightarrow> bool"
   354   where
   355 "differentiate_equality f_f v_v = (
   356   let
   357     f_f' = Take ((primed (lhs f_f)) = d_d v_v (rhs f_f))
   358   in (
   359     (Try (Rewrite_Set_Inst [(''bdv'',v_v)] ''diff_conv'' )) #> (
   360     Repeat (
   361       (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''diff_sum'')) Or
   362       (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''diff_dif''        )) Or
   363       (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''diff_prod_const'')) Or
   364       (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''diff_prod'')) Or
   365       (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''diff_quot'')) Or
   366       (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''diff_sin'')) Or
   367       (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''diff_sin_chain'')) Or
   368       (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''diff_cos'')) Or
   369       (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''diff_cos_chain'')) Or
   370       (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''diff_pow'')) Or
   371       (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''diff_pow_chain'')) Or
   372       (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''diff_ln'')) Or
   373       (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''diff_ln_chain'')) Or
   374       (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''diff_exp'')) Or
   375       (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''diff_exp_chain'')) Or
   376       (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''diff_const'')) Or
   377       (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''diff_var'')) Or
   378       (Repeat (Rewrite_Set ''make_polynomial'')))) #> (
   379     Try (Rewrite_Set_Inst [(''bdv'', v_v)] ''diff_sym_conv'' ))
   380     ) f_f')"
   381 setup \<open>KEStore_Elems.add_mets
   382     [MethodC.prep_input @{theory} "met_diff_equ" [] MethodC.id_empty
   383       (["diff", "differentiate_equality"],
   384         [("#Given" ,["functionEq f_f", "differentiateFor v_v"]),
   385           ("#Find"  ,["derivativeEq f_f'"])],
   386         {rew_ord'="tless_true", rls' = erls_diff, calc = [], srls = srls_diff, prls=Rule_Set.empty,
   387           crls=Atools_erls, errpats = [], nrls = norm_diff},
   388         @{thm differentiate_equality.simps})]
   389 \<close>
   390 
   391 partial_function (tailrec) simplify_derivative :: "real \<Rightarrow> real \<Rightarrow> real"
   392   where
   393 "simplify_derivative term bound_variable = (
   394   let
   395    term' = Take (d_d bound_variable term)
   396   in (
   397     (Try (Rewrite_Set ''norm_Rational'')) #>
   398     (Try (Rewrite_Set_Inst [(''bdv'', bound_variable)] ''diff_conv'')) #>
   399     (Try (Rewrite_Set_Inst [(''bdv'', bound_variable)] ''norm_diff'')) #>
   400     (Try (Rewrite_Set_Inst [(''bdv'', bound_variable)] ''diff_sym_conv'')) #>
   401     (Try (Rewrite_Set ''norm_Rational''))
   402     ) term')"
   403 
   404 setup \<open>KEStore_Elems.add_mets
   405     [MethodC.prep_input @{theory} "met_diff_after_simp" [] MethodC.id_empty
   406       (["diff", "after_simplification"],
   407         [("#Given" ,["functionTerm term", "differentiateFor bound_variable"]),
   408           ("#Find"  ,["derivative term'"])],
   409         {rew_ord'="tless_true", rls' = Rule_Set.empty, calc = [], srls = Rule_Set.empty, prls=Rule_Set.empty,
   410           crls=Atools_erls, errpats = [], nrls = norm_Rational},
   411         @{thm simplify_derivative.simps})]
   412 \<close>
   413 setup \<open>KEStore_Elems.add_cas
   414   [((Thm.term_of o the o (TermC.parse @{theory})) "Diff",
   415 	      (("Isac_Knowledge", ["derivative_of", "function"], ["no_met"]), argl2dtss))]\<close>
   416 ML \<open>
   417 
   418 (*.handle cas-input like "Differentiate (A = s * (a - s), s)".*)
   419 (* val (t, pairl) = strip_comb (str2term "Differentiate (A = s * (a - s), s)");
   420    val [Const ("Product_Type.Pair", _) $ t $ bdv] = pairl;
   421    *)
   422 fun argl2dtss [Const ("Product_Type.Pair", _) $ t $ bdv] =
   423     [((Thm.term_of o the o (TermC.parse \<^theory>)) "functionEq", [t]),
   424      ((Thm.term_of o the o (TermC.parse \<^theory>)) "differentiateFor", [bdv]),
   425      ((Thm.term_of o the o (TermC.parse \<^theory>)) "derivativeEq", 
   426       [(Thm.term_of o the o (TermC.parse \<^theory>)) "f_f'::bool"])
   427      ]
   428   | argl2dtss _ = raise ERROR "Diff.ML: wrong argument for argl2dtss";
   429 \<close>
   430 setup \<open>KEStore_Elems.add_cas
   431   [((Thm.term_of o the o (TermC.parse @{theory})) "Differentiate",  
   432 	      (("Isac_Knowledge", ["named", "derivative_of", "function"], ["no_met"]), argl2dtss))]
   433 \<close> ML \<open>
   434 \<close> ML \<open>
   435 \<close>
   436 end