src/Tools/isac/Knowledge/Diff.thy
author wenzelm
Thu, 10 Jun 2021 11:54:20 +0200
changeset 60289 a7b88fc19a92
parent 60286 31efa1b39a20
child 60290 bb4e8b01b072
permissions -rw-r--r--
clarified command name: this is to register already defined rule sets in the Knowledge Store;
     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 val thy = @{theory};
    91 
    92 (** eval functions **)
    93 
    94 fun primed (Const (id, T)) = Const (id ^ "'", T)
    95   | primed (Free (id, T)) = Free (id ^ "'", T)
    96   | primed t = raise ERROR ("primed called with arg = '"^ UnparseC.term t ^"'");
    97 
    98 (*("primed", ("Diff.primed", eval_primed "#primed"))*)
    99 fun eval_primed _ _ (p as (Const ("Diff.primed",_) $ t)) _ =
   100     SOME ((UnparseC.term p) ^ " = " ^ UnparseC.term (primed t),
   101 	  HOLogic.Trueprop $ (TermC.mk_equality (p, primed t)))
   102   | eval_primed _ _ _ _ = NONE;
   103 \<close>
   104 setup \<open>KEStore_Elems.add_calcs
   105   [("primed", ("Diff.primed", eval_primed "#primed"))]\<close>
   106 ML \<open>
   107 (** rulesets **)
   108 
   109 (*.converts a term such that differentiation works optimally.*)
   110 val diff_conv =   
   111     Rule_Def.Repeat {id="diff_conv", 
   112 	 preconds = [], 
   113 	 rew_ord = ("termlessI",termlessI), 
   114 	 erls = Rule_Set.append_rules "erls_diff_conv" Rule_Set.empty 
   115 			   [Rule.Eval ("Prog_Expr.occurs_in", Prog_Expr.eval_occurs_in ""),
   116 			    Rule.Thm ("not_true",ThmC.numerals_to_Free @{thm not_true}),
   117 			    Rule.Thm ("not_false",ThmC.numerals_to_Free @{thm not_false}),
   118 			    Rule.Eval ("Orderings.ord_class.less", Prog_Expr.eval_equ "#less_"),
   119 			    Rule.Thm ("and_true",ThmC.numerals_to_Free @{thm and_true}),
   120 			    Rule.Thm ("and_false",ThmC.numerals_to_Free @{thm and_false})
   121 			    ], 
   122 	 srls = Rule_Set.Empty, calc = [], errpatts = [],
   123 	 rules =
   124   [Rule.Thm ("frac_conv", ThmC.numerals_to_Free @{thm frac_conv}),
   125      (*"?bdv occurs_in ?b \<Longrightarrow> 0 < ?n \<Longrightarrow> ?a / ?b \<up> ?n = ?a * ?b \<up> - ?n"*)
   126 		   Rule.Thm ("sqrt_conv_bdv", ThmC.numerals_to_Free @{thm sqrt_conv_bdv}),
   127 		     (*"sqrt ?bdv = ?bdv \<up> (1 / 2)"*)
   128 		   Rule.Thm ("sqrt_conv_bdv_n", ThmC.numerals_to_Free @{thm sqrt_conv_bdv_n}),
   129 		     (*"sqrt (?bdv \<up> ?n) = ?bdv \<up> (?n / 2)"*)
   130 		   Rule.Thm ("sqrt_conv", ThmC.numerals_to_Free @{thm sqrt_conv}),
   131 		     (*"?bdv occurs_in ?u \<Longrightarrow> sqrt ?u = ?u \<up> (1 / 2)"*)
   132 		   Rule.Thm ("root_conv", ThmC.numerals_to_Free @{thm root_conv}),
   133 		     (*"?bdv occurs_in ?u \<Longrightarrow> nroot ?n ?u = ?u \<up> (1 / ?n)"*)
   134 		   Rule.Thm ("realpow_pow_bdv", ThmC.numerals_to_Free @{thm realpow_pow_bdv}),
   135 		     (* "(?bdv \<up> ?b) \<up> ?c = ?bdv \<up> (?b * ?c)"*)
   136 		   Rule.Eval ("Groups.times_class.times", (**)eval_binop "#mult_"),
   137 		   Rule.Thm ("rat_mult",ThmC.numerals_to_Free @{thm rat_mult}),
   138 		     (*a / b * (c / d) = a * c / (b * d)*)
   139 		   Rule.Thm ("times_divide_eq_right",ThmC.numerals_to_Free @{thm times_divide_eq_right}),
   140 		     (*?x * (?y / ?z) = ?x * ?y / ?z*)
   141 		   Rule.Thm ("times_divide_eq_left",ThmC.numerals_to_Free @{thm times_divide_eq_left})
   142 		     (*?y / ?z * ?x = ?y * ?x / ?z*)
   143 		 ],
   144 	 scr = Rule.Empty_Prog};
   145 \<close>
   146 ML \<open>
   147 (*.beautifies a term after differentiation.*)
   148 val diff_sym_conv =   
   149     Rule_Def.Repeat {id="diff_sym_conv", 
   150 	 preconds = [], 
   151 	 rew_ord = ("termlessI",termlessI), 
   152 	 erls = Rule_Set.append_rules "erls_diff_sym_conv" Rule_Set.empty 
   153 			   [Rule.Eval ("Orderings.ord_class.less", Prog_Expr.eval_equ "#less_")
   154 			    ], 
   155 	 srls = Rule_Set.Empty, calc = [], errpatts = [],
   156 	 rules = [Rule.Thm ("frac_sym_conv", ThmC.numerals_to_Free @{thm frac_sym_conv}),
   157 		  Rule.Thm ("sqrt_sym_conv", ThmC.numerals_to_Free @{thm sqrt_sym_conv}),
   158 		  Rule.Thm ("root_sym_conv", ThmC.numerals_to_Free @{thm root_sym_conv}),
   159 		  Rule.Thm ("sym_real_mult_minus1",
   160 		       ThmC.numerals_to_Free (@{thm real_mult_minus1} RS @{thm sym})),
   161 		      (*- ?z = "-1 * ?z"*)
   162 		  Rule.Thm ("rat_mult",ThmC.numerals_to_Free @{thm rat_mult}),
   163 		  (*a / b * (c / d) = a * c / (b * d)*)
   164 		  Rule.Thm ("times_divide_eq_right",ThmC.numerals_to_Free @{thm times_divide_eq_right}),
   165 		  (*?x * (?y / ?z) = ?x * ?y / ?z*)
   166 		  Rule.Thm ("times_divide_eq_left",ThmC.numerals_to_Free @{thm times_divide_eq_left}),
   167 		  (*?y / ?z * ?x = ?y * ?x / ?z*)
   168 		  Rule.Eval ("Groups.times_class.times", (**)eval_binop "#mult_")
   169 		 ],
   170 	 scr = Rule.Empty_Prog};
   171 
   172 (*..*)
   173 val srls_diff = 
   174     Rule_Def.Repeat {id="srls_differentiate..", 
   175 	 preconds = [], 
   176 	 rew_ord = ("termlessI",termlessI), 
   177 	 erls = Rule_Set.empty, 
   178 	 srls = Rule_Set.Empty, calc = [], errpatts = [],
   179 	 rules = [Rule.Eval("Prog_Expr.lhs", Prog_Expr.eval_lhs "eval_lhs_"),
   180 		  Rule.Eval("Prog_Expr.rhs", Prog_Expr.eval_rhs "eval_rhs_"),
   181 		  Rule.Eval("Diff.primed", eval_primed "Diff.primed")
   182 		  ],
   183 	 scr = Rule.Empty_Prog};
   184 \<close>
   185 ML \<open>
   186 (*..*)
   187 val erls_diff = 
   188     Rule_Set.append_rules "erls_differentiate.." Rule_Set.empty
   189                [Rule.Thm ("not_true",ThmC.numerals_to_Free @{thm not_true}),
   190 		Rule.Thm ("not_false",ThmC.numerals_to_Free @{thm not_false}),
   191 		
   192 		Rule.Eval ("Prog_Expr.ident", Prog_Expr.eval_ident "#ident_"),    
   193 		Rule.Eval ("Prog_Expr.is_atom", Prog_Expr.eval_is_atom "#is_atom_"),
   194 		Rule.Eval ("Prog_Expr.occurs_in", Prog_Expr.eval_occurs_in ""),
   195 		Rule.Eval ("Prog_Expr.is_const", Prog_Expr.eval_const "#is_const_")
   196 		];
   197 
   198 (*.rules for differentiation, _no_ simplification.*)
   199 val diff_rules =
   200     Rule_Def.Repeat {id="diff_rules", preconds = [], rew_ord = ("termlessI",termlessI), 
   201 	 erls = erls_diff, srls = Rule_Set.Empty, calc = [], errpatts = [],
   202 	 rules = [Rule.Thm ("diff_sum",ThmC.numerals_to_Free @{thm diff_sum}),
   203 		  Rule.Thm ("diff_dif",ThmC.numerals_to_Free @{thm diff_dif}),
   204 		  Rule.Thm ("diff_prod_const",ThmC.numerals_to_Free @{thm diff_prod_const}),
   205 		  Rule.Thm ("diff_prod",ThmC.numerals_to_Free @{thm diff_prod}),
   206 		  Rule.Thm ("diff_quot",ThmC.numerals_to_Free @{thm diff_quot}),
   207 		  Rule.Thm ("diff_sin",ThmC.numerals_to_Free @{thm diff_sin}),
   208 		  Rule.Thm ("diff_sin_chain",ThmC.numerals_to_Free @{thm diff_sin_chain}),
   209 		  Rule.Thm ("diff_cos",ThmC.numerals_to_Free @{thm diff_cos}),
   210 		  Rule.Thm ("diff_cos_chain",ThmC.numerals_to_Free @{thm diff_cos_chain}),
   211 		  Rule.Thm ("diff_pow",ThmC.numerals_to_Free @{thm diff_pow}),
   212 		  Rule.Thm ("diff_pow_chain",ThmC.numerals_to_Free @{thm diff_pow_chain}),
   213 		  Rule.Thm ("diff_ln",ThmC.numerals_to_Free @{thm diff_ln}),
   214 		  Rule.Thm ("diff_ln_chain",ThmC.numerals_to_Free @{thm diff_ln_chain}),
   215 		  Rule.Thm ("diff_exp",ThmC.numerals_to_Free @{thm diff_exp}),
   216 		  Rule.Thm ("diff_exp_chain",ThmC.numerals_to_Free @{thm diff_exp_chain}),
   217 (*
   218 		  Rule.Thm ("diff_sqrt",ThmC.numerals_to_Free @{thm diff_sqrt}),
   219 		  Rule.Thm ("diff_sqrt_chain",ThmC.numerals_to_Free @{thm diff_sqrt_chain}),
   220 *)
   221 		  Rule.Thm ("diff_const",ThmC.numerals_to_Free @{thm diff_const}),
   222 		  Rule.Thm ("diff_var",ThmC.numerals_to_Free @{thm diff_var})
   223 		  ],
   224 	 scr = Rule.Empty_Prog};
   225 \<close>
   226 ML \<open>
   227 (*.normalisation for checking user-input.*)
   228 val norm_diff = 
   229   Rule_Def.Repeat
   230     {id="norm_diff", preconds = [], rew_ord = ("termlessI",termlessI), 
   231      erls = Rule_Set.Empty, srls = Rule_Set.Empty, calc = [], errpatts = [],
   232      rules = [Rule.Rls_ diff_rules, Rule.Rls_ norm_Poly ],
   233      scr = Rule.Empty_Prog};
   234 \<close>
   235 rule_set_knowledge
   236   erls_diff = \<open>prep_rls' erls_diff\<close> and
   237   diff_rules = \<open>prep_rls' diff_rules\<close> and
   238   norm_diff = \<open>prep_rls' norm_diff\<close> and
   239   diff_conv = \<open>prep_rls' diff_conv\<close> and
   240   diff_sym_conv = \<open>prep_rls' diff_sym_conv\<close>
   241 
   242 (** problem types **)
   243 setup \<open>KEStore_Elems.add_pbts
   244   [(Problem.prep_input thy "pbl_fun" [] Problem.id_empty (["function"], [], Rule_Set.empty, NONE, [])),
   245     (Problem.prep_input thy "pbl_fun_deriv" [] Problem.id_empty
   246       (["derivative_of", "function"],
   247         [("#Given" ,["functionTerm f_f", "differentiateFor v_v"]),
   248           ("#Find"  ,["derivative f_f'"])],
   249         Rule_Set.append_rules "empty" Rule_Set.empty [],
   250         SOME "Diff (f_f, v_v)", [["diff", "differentiate_on_R"],
   251 			  ["diff", "after_simplification"]])),
   252     (*here "named" is used differently from Integration"*)
   253     (Problem.prep_input thy "pbl_fun_deriv_nam" [] Problem.id_empty
   254       (["named", "derivative_of", "function"],
   255         [("#Given" ,["functionEq f_f", "differentiateFor v_v"]),
   256           ("#Find"  ,["derivativeEq f_f'"])],
   257         Rule_Set.append_rules "empty" Rule_Set.empty [],
   258         SOME "Differentiate (f_f, v_v)",
   259         [["diff", "differentiate_equality"]]))]\<close>
   260 
   261 ML \<open>
   262 (** CAS-commands **)
   263 
   264 (*.handle cas-input like "Diff (a * x^3 + b, x)".*)
   265 (* val (t, pairl) = strip_comb (str2term "Diff (a * x^3 + b, x)");
   266    val [Const ("Product_Type.Pair", _) $ t $ bdv] = pairl;
   267    *)
   268 fun argl2dtss [Const ("Product_Type.Pair", _) $ t $ bdv] =
   269     [((Thm.term_of o the o (TermC.parse thy)) "functionTerm", [t]),
   270      ((Thm.term_of o the o (TermC.parse thy)) "differentiateFor", [bdv]),
   271      ((Thm.term_of o the o (TermC.parse thy)) "derivative", 
   272       [(Thm.term_of o the o (TermC.parse thy)) "f_f'"])
   273      ]
   274   | argl2dtss _ = raise ERROR "Diff.ML: wrong argument for argl2dtss";
   275 \<close>
   276 setup \<open>KEStore_Elems.add_mets
   277     [MethodC.prep_input thy "met_diff" [] MethodC.id_empty
   278       (["diff"], [],
   279         {rew_ord'="tless_true",rls'=Atools_erls,calc = [], srls = Rule_Set.empty, prls=Rule_Set.empty,
   280           crls = Atools_erls, errpats = [], nrls = norm_diff},
   281         @{thm refl})]
   282 \<close>
   283 
   284 partial_function (tailrec) differentiate_on_R :: "real \<Rightarrow> real \<Rightarrow> real"
   285   where
   286 "differentiate_on_R f_f v_v = (
   287   let
   288     f_f' = Take (d_d v_v f_f)
   289   in (
   290     (Try (Rewrite_Set_Inst [(''bdv'',v_v)] ''diff_conv'')) #> (
   291     Repeat (
   292       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_sum'')) Or
   293       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_prod_const'')) Or
   294       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_prod'')) Or
   295       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_quot'')) Or
   296       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_sin'')) Or
   297       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_sin_chain'')) Or
   298       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_cos'')) Or
   299       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_cos_chain'')) Or
   300       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_pow'')) Or
   301       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_pow_chain'')) Or
   302       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_ln'')) Or
   303       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_ln_chain'')) Or
   304       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_exp'')) Or
   305       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_exp_chain'')) Or
   306       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_const'')) Or
   307       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_var'')) Or
   308       (Repeat (Rewrite_Set ''make_polynomial'')))) #> (
   309     Try (Rewrite_Set_Inst [(''bdv'',v_v)] ''diff_sym_conv''))
   310     ) f_f')"
   311 setup \<open>KEStore_Elems.add_mets
   312     [MethodC.prep_input thy "met_diff_onR" [] MethodC.id_empty
   313       (["diff", "differentiate_on_R"],
   314         [("#Given" ,["functionTerm f_f", "differentiateFor v_v"]),
   315           ("#Find"  ,["derivative f_f'"])],
   316         {rew_ord'="tless_true", rls' = erls_diff, calc = [], srls = Rule_Set.empty, prls=Rule_Set.empty,
   317           crls = Atools_erls, errpats = [], nrls = norm_diff},
   318         @{thm differentiate_on_R.simps})]
   319 \<close>
   320 
   321 partial_function (tailrec) differentiateX :: "real \<Rightarrow> real \<Rightarrow> real"
   322   where
   323 "differentiateX f_f v_v = (
   324   let
   325     f_f' = Take (d_d v_v f_f)
   326   in (
   327     Repeat (
   328       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_sum'')) Or
   329       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_prod_const'' )) Or
   330       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_prod'')) Or
   331       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_quot'')) Or
   332       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_sin'')) Or
   333       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_sin_chain'')) Or
   334       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_cos'')) Or
   335       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_cos_chain'')) Or
   336       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_pow'')) Or
   337       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_pow_chain'')) Or
   338       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_ln'')) Or
   339       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_ln_chain'')) Or
   340       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_exp'')) Or
   341       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_exp_chain'')) Or
   342       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_const'')) Or
   343       (Repeat (Rewrite_Inst [(''bdv'',v_v)] ''diff_var'')) Or
   344       (Repeat (Rewrite_Set ''make_polynomial'')))
   345     ) f_f')"
   346 setup \<open>KEStore_Elems.add_mets
   347     [MethodC.prep_input thy "met_diff_simpl" [] MethodC.id_empty
   348       (["diff", "diff_simpl"],
   349         [("#Given", ["functionTerm f_f", "differentiateFor v_v"]),
   350          ("#Find" , ["derivative f_f'"])],
   351         {rew_ord'="tless_true", rls' = erls_diff, calc = [], srls = Rule_Set.empty, prls=Rule_Set.empty,
   352           crls = Atools_erls, errpats = [], nrls = norm_diff},
   353         @{thm differentiateX.simps})]
   354 \<close>
   355 
   356 partial_function (tailrec) differentiate_equality :: "bool \<Rightarrow> real \<Rightarrow> bool"
   357   where
   358 "differentiate_equality f_f v_v = (
   359   let
   360     f_f' = Take ((primed (lhs f_f)) = d_d v_v (rhs f_f))
   361   in (
   362     (Try (Rewrite_Set_Inst [(''bdv'',v_v)] ''diff_conv'' )) #> (
   363     Repeat (
   364       (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''diff_sum'')) Or
   365       (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''diff_dif''        )) Or
   366       (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''diff_prod_const'')) Or
   367       (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''diff_prod'')) Or
   368       (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''diff_quot'')) Or
   369       (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''diff_sin'')) Or
   370       (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''diff_sin_chain'')) Or
   371       (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''diff_cos'')) Or
   372       (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''diff_cos_chain'')) Or
   373       (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''diff_pow'')) Or
   374       (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''diff_pow_chain'')) Or
   375       (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''diff_ln'')) Or
   376       (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''diff_ln_chain'')) Or
   377       (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''diff_exp'')) Or
   378       (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''diff_exp_chain'')) Or
   379       (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''diff_const'')) Or
   380       (Repeat (Rewrite_Inst [(''bdv'', v_v)] ''diff_var'')) Or
   381       (Repeat (Rewrite_Set ''make_polynomial'')))) #> (
   382     Try (Rewrite_Set_Inst [(''bdv'', v_v)] ''diff_sym_conv'' ))
   383     ) f_f')"
   384 setup \<open>KEStore_Elems.add_mets
   385     [MethodC.prep_input thy "met_diff_equ" [] MethodC.id_empty
   386       (["diff", "differentiate_equality"],
   387         [("#Given" ,["functionEq f_f", "differentiateFor v_v"]),
   388           ("#Find"  ,["derivativeEq f_f'"])],
   389         {rew_ord'="tless_true", rls' = erls_diff, calc = [], srls = srls_diff, prls=Rule_Set.empty,
   390           crls=Atools_erls, errpats = [], nrls = norm_diff},
   391         @{thm differentiate_equality.simps})]
   392 \<close>
   393 
   394 partial_function (tailrec) simplify_derivative :: "real \<Rightarrow> real \<Rightarrow> real"
   395   where
   396 "simplify_derivative term bound_variable = (
   397   let
   398    term' = Take (d_d bound_variable term)
   399   in (
   400     (Try (Rewrite_Set ''norm_Rational'')) #>
   401     (Try (Rewrite_Set_Inst [(''bdv'', bound_variable)] ''diff_conv'')) #>
   402     (Try (Rewrite_Set_Inst [(''bdv'', bound_variable)] ''norm_diff'')) #>
   403     (Try (Rewrite_Set_Inst [(''bdv'', bound_variable)] ''diff_sym_conv'')) #>
   404     (Try (Rewrite_Set ''norm_Rational''))
   405     ) term')"
   406 
   407 setup \<open>KEStore_Elems.add_mets
   408     [MethodC.prep_input thy "met_diff_after_simp" [] MethodC.id_empty
   409       (["diff", "after_simplification"],
   410         [("#Given" ,["functionTerm term", "differentiateFor bound_variable"]),
   411           ("#Find"  ,["derivative term'"])],
   412         {rew_ord'="tless_true", rls' = Rule_Set.empty, calc = [], srls = Rule_Set.empty, prls=Rule_Set.empty,
   413           crls=Atools_erls, errpats = [], nrls = norm_Rational},
   414         @{thm simplify_derivative.simps})]
   415 \<close>
   416 setup \<open>KEStore_Elems.add_cas
   417   [((Thm.term_of o the o (TermC.parse thy)) "Diff",
   418 	      (("Isac_Knowledge", ["derivative_of", "function"], ["no_met"]), argl2dtss))]\<close>
   419 ML \<open>
   420 
   421 (*.handle cas-input like "Differentiate (A = s * (a - s), s)".*)
   422 (* val (t, pairl) = strip_comb (str2term "Differentiate (A = s * (a - s), s)");
   423    val [Const ("Product_Type.Pair", _) $ t $ bdv] = pairl;
   424    *)
   425 fun argl2dtss [Const ("Product_Type.Pair", _) $ t $ bdv] =
   426     [((Thm.term_of o the o (TermC.parse thy)) "functionEq", [t]),
   427      ((Thm.term_of o the o (TermC.parse thy)) "differentiateFor", [bdv]),
   428      ((Thm.term_of o the o (TermC.parse thy)) "derivativeEq", 
   429       [(Thm.term_of o the o (TermC.parse thy)) "f_f'::bool"])
   430      ]
   431   | argl2dtss _ = raise ERROR "Diff.ML: wrong argument for argl2dtss";
   432 \<close>
   433 setup \<open>KEStore_Elems.add_cas
   434   [((Thm.term_of o the o (TermC.parse thy)) "Differentiate",  
   435 	      (("Isac_Knowledge", ["named", "derivative_of", "function"], ["no_met"]), argl2dtss))]
   436 \<close> ML \<open>
   437 \<close> ML \<open>
   438 \<close>
   439 end