src/Tools/isac/Knowledge/Integrate.thy
author wenzelm
Sun, 20 Jun 2021 16:26:18 +0200
changeset 60306 51ec2e101e9f
parent 60303 815b0dc8b589
child 60309 70a1d102660d
permissions -rw-r--r--
Isar command 'problem' as combination of KEStore_Elems.add_pbts + Problem.prep_import, without change of semantics;
     1 (* integration over the reals
     2    author: Walther Neuper
     3    050814, 08:51
     4    (c) due to copyright terms
     5 *)
     6 
     7 theory Integrate imports Diff begin
     8 
     9 consts
    10 
    11   Integral            :: "[real, real]=> real" ("Integral _ D _" 91)
    12 (*new_c	      :: "real => real"        ("new_c _" 66)*)
    13   is_f_x            :: "real => bool"        ("_ is'_f'_x" 10)
    14 
    15   (*descriptions in the related problems*)
    16   integrateBy         :: "real => una"
    17   antiDerivative      :: "real => una"
    18   antiDerivativeName  :: "(real => real) => una"
    19 
    20   (*the CAS-command, eg. "Integrate (2*x \<up> 3, x)"*)
    21   Integrate           :: "[real * real] => real"
    22 
    23 axiomatization where
    24 (*stated as axioms, todo: prove as theorems
    25   'bdv' is a constant handled on the meta-level 
    26    specifically as a 'bound variable'            *)
    27 
    28 (*Ambiguous input\<^here> produces 3 parse trees -----------------------------\\*)
    29   integral_const:    "Not (bdv occurs_in u) ==> Integral u D bdv = u * bdv" and
    30   integral_var:      "Integral bdv D bdv = bdv \<up> 2 / 2" and
    31 
    32   integral_add:      "Integral (u + v) D bdv =  
    33 		     (Integral u D bdv) + (Integral v D bdv)" and
    34   integral_mult:     "[| Not (bdv occurs_in u); bdv occurs_in v |] ==>  
    35 		     Integral (u * v) D bdv = u * (Integral v D bdv)" and
    36 (*WN080222: this goes into sub-terms, too ...
    37   call_for_new_c:    "[| Not (matches (u + new_c v) a); Not (a is_f_x) |] ==>  
    38 		     a = a + new_c a"
    39 *)
    40   integral_pow:      "Integral bdv \<up> n D bdv = bdv \<up> (n+1) / (n + 1)"
    41 (*Ambiguous input\<^here> produces 3 parse trees -----------------------------//*)
    42 
    43 ML \<open>
    44 (** eval functions **)
    45 
    46 val c = Free ("c", HOLogic.realT);
    47 (*.create a new unique variable 'c..' in a term; for use by Rule.Eval in a rls;
    48    an alternative to do this would be '(Try (Calculate new_c_) (new_c es__))'
    49    in the script; this will be possible if currying doesnt take the value
    50    from a variable, but the value '(new_c es__)' itself.*)
    51 fun new_c term = 
    52     let fun selc var = 
    53 	    case (Symbol.explode o id_of) var of
    54 		"c"::[] => true
    55 	      |	"c"::"_"::is => (case (TermC.int_opt_of_string o implode) is of
    56 				     SOME _ => true
    57 				   | NONE => false)
    58               | _ => false;
    59 	fun get_coeff c = case (Symbol.explode o id_of) c of
    60 	      		      "c"::"_"::is => (the o TermC.int_opt_of_string o implode) is
    61 			    | _ => 0;
    62         val cs = filter selc (TermC.vars term);
    63     in 
    64 	case cs of
    65 	    [] => c
    66 	  | [_] => Free ("c_2", HOLogic.realT)
    67 	  | cs => 
    68 	    let val max_coeff = maxl (map get_coeff cs)
    69 	    in Free ("c_"^string_of_int (max_coeff + 1), HOLogic.realT) end
    70     end;
    71 
    72 (*WN080222
    73 (*("new_c", ("Integrate.new_c", eval_new_c "#new_c_"))*)
    74 fun eval_new_c _ _ (p as (Const ("Integrate.new_c",_) $ t)) _ =
    75      SOME ((UnparseC.term p) ^ " = " ^ UnparseC.term (new_c p),
    76 	  Trueprop $ (mk_equality (p, new_c p)))
    77   | eval_new_c _ _ _ _ = NONE;
    78 *)
    79 
    80 (*WN080222:*)
    81 (*("add_new_c", ("Integrate.add_new_c", eval_add_new_c "#add_new_c_"))
    82   add a new c to a term or a fun-equation;
    83   this is _not in_ the term, because only applied to _whole_ term*)
    84 fun eval_add_new_c (_:string) "Integrate.add_new_c" p (_:theory) =
    85     let val p' = case p of
    86 		     Const ("HOL.eq", T) $ lh $ rh => 
    87 		     Const ("HOL.eq", T) $ lh $ TermC.mk_add rh (new_c rh)
    88 		   | p => TermC.mk_add p (new_c p)
    89     in SOME ((UnparseC.term p) ^ " = " ^ UnparseC.term p',
    90 	  HOLogic.Trueprop $ (TermC.mk_equality (p, p')))
    91     end
    92   | eval_add_new_c _ _ _ _ = NONE;
    93 
    94 
    95 (*("is_f_x", ("Integrate.is_f_x", eval_is_f_x "is_f_x_"))*)
    96 fun eval_is_f_x _ _(p as (Const ("Integrate.is_f_x", _)
    97 					   $ arg)) _ =
    98     if TermC.is_f_x arg
    99     then SOME ((UnparseC.term p) ^ " = True",
   100 	       HOLogic.Trueprop $ (TermC.mk_equality (p, @{term True})))
   101     else SOME ((UnparseC.term p) ^ " = False",
   102 	       HOLogic.Trueprop $ (TermC.mk_equality (p, @{term False})))
   103   | eval_is_f_x _ _ _ _ = NONE;
   104 \<close>
   105 setup \<open>KEStore_Elems.add_calcs
   106   [("add_new_c", ("Integrate.add_new_c", eval_add_new_c "add_new_c_")),
   107     ("is_f_x", ("Integrate.is_f_x", eval_is_f_x "is_f_idextifier_"))]\<close>
   108 ML \<open>
   109 (** rulesets **)
   110 
   111 (*.rulesets for integration.*)
   112 val integration_rules = 
   113     Rule_Def.Repeat {id="integration_rules", preconds = [], 
   114 	 rew_ord = ("termlessI",termlessI), 
   115 	 erls = Rule_Def.Repeat {id="conditions_in_integration_rules", 
   116 		     preconds = [], 
   117 		     rew_ord = ("termlessI",termlessI), 
   118 		     erls = Rule_Set.Empty, 
   119 		     srls = Rule_Set.Empty, calc = [], errpatts = [],
   120 		     rules = [(*for rewriting conditions in Thm's*)
   121 			      \<^rule_eval>\<open>Prog_Expr.occurs_in\<close> (Prog_Expr.eval_occurs_in "#occurs_in_"),
   122 			      \<^rule_thm>\<open>not_true\<close>,
   123 			      \<^rule_thm>\<open>not_false\<close>
   124 			      ],
   125 		     scr = Rule.Empty_Prog}, 
   126 	 srls = Rule_Set.Empty, calc = [], errpatts = [],
   127 	 rules = [
   128 		  \<^rule_thm>\<open>integral_const\<close>,
   129 		  \<^rule_thm>\<open>integral_var\<close>,
   130 		  \<^rule_thm>\<open>integral_add\<close>,
   131 		  \<^rule_thm>\<open>integral_mult\<close>,
   132 		  \<^rule_thm>\<open>integral_pow\<close>,
   133 		  \<^rule_eval>\<open>plus\<close> (**)(eval_binop "#add_")(*for n+1*)
   134 		  ],
   135 	 scr = Rule.Empty_Prog};
   136 \<close>
   137 ML \<open>
   138 val add_new_c = 
   139     Rule_Set.Sequence {id="add_new_c", preconds = [], 
   140 	 rew_ord = ("termlessI",termlessI), 
   141 	 erls = Rule_Def.Repeat {id="conditions_in_add_new_c", 
   142 		     preconds = [], 
   143 		     rew_ord = ("termlessI",termlessI), 
   144 		     erls = Rule_Set.Empty, 
   145 		     srls = Rule_Set.Empty, calc = [], errpatts = [],
   146 		     rules = [\<^rule_eval>\<open>Prog_Expr.matches\<close> (Prog_Expr.eval_matches""),
   147 			      \<^rule_eval>\<open>Integrate.is_f_x\<close> (eval_is_f_x "is_f_x_"),
   148 			      \<^rule_thm>\<open>not_true\<close>,
   149 			      \<^rule_thm>\<open>not_false\<close>
   150 			      ],
   151 		     scr = Rule.Empty_Prog}, 
   152 	 srls = Rule_Set.Empty, calc = [], errpatts = [],
   153 	 rules = [ (*\<^rule_thm>\<open>call_for_new_c\<close>,*)
   154 		   Rule.Cal1 ("Integrate.add_new_c", eval_add_new_c "new_c_")
   155 		   ],
   156 	 scr = Rule.Empty_Prog};
   157 \<close>
   158 ML \<open>
   159 
   160 (*.rulesets for simplifying Integrals.*)
   161 
   162 (*.for simplify_Integral adapted from 'norm_Rational_rls'.*)
   163 val norm_Rational_rls_noadd_fractions = 
   164 Rule_Def.Repeat {id = "norm_Rational_rls_noadd_fractions", preconds = [], 
   165      rew_ord = ("dummy_ord",Rewrite_Ord.dummy_ord), 
   166      erls = norm_rat_erls, srls = Rule_Set.Empty, calc = [], errpatts = [],
   167      rules = [(*Rule.Rls_ add_fractions_p_rls,!!!*)
   168 	      Rule.Rls_ (*rat_mult_div_pow original corrected WN051028*)
   169 		  (Rule_Def.Repeat {id = "rat_mult_div_pow", preconds = [], 
   170 		       rew_ord = ("dummy_ord",Rewrite_Ord.dummy_ord), 
   171 		       erls = (*FIXME.WN051028 Rule_Set.empty,*)
   172 		       Rule_Set.append_rules "Rule_Set.empty-is_polyexp" Rule_Set.empty
   173 				  [\<^rule_eval>\<open>is_polyexp\<close> (eval_is_polyexp "")],
   174 				  srls = Rule_Set.Empty, calc = [], errpatts = [],
   175 				  rules = [\<^rule_thm>\<open>rat_mult\<close>,
   176 	       (*"?a / ?b * (?c / ?d) = ?a * ?c / (?b * ?d)"*)
   177 	       \<^rule_thm>\<open>rat_mult_poly_l\<close>,
   178 	       (*"?c is_polyexp ==> ?c * (?a / ?b) = ?c * ?a / ?b"*)
   179 	       \<^rule_thm>\<open>rat_mult_poly_r\<close>,
   180 	       (*"?c is_polyexp ==> ?a / ?b * ?c = ?a * ?c / ?b"*)
   181 
   182 	       \<^rule_thm>\<open>real_divide_divide1_mg\<close>,
   183 	       (*"y ~= 0 ==> (u / v) / (y / z) = (u * z) / (y * v)"*)
   184 	       \<^rule_thm>\<open>divide_divide_eq_right\<close>,
   185 	       (*"?x / (?y / ?z) = ?x * ?z / ?y"*)
   186 	       \<^rule_thm>\<open>divide_divide_eq_left\<close>,
   187 	       (*"?x / ?y / ?z = ?x / (?y * ?z)"*)
   188 	       \<^rule_eval>\<open>divide\<close> (Prog_Expr.eval_cancel "#divide_e"),
   189 	      
   190 	       \<^rule_thm>\<open>rat_power\<close>
   191 		(*"(?a / ?b)  \<up>  ?n = ?a  \<up>  ?n / ?b  \<up>  ?n"*)
   192 	       ],
   193       scr = Rule.Empty_Prog
   194       }),
   195 		Rule.Rls_ make_rat_poly_with_parentheses,
   196 		Rule.Rls_ cancel_p_rls,(*FIXME:cancel_p does NOT order sometimes*)
   197 		Rule.Rls_ rat_reduce_1
   198 		],
   199        scr = Rule.Empty_Prog
   200        };
   201 
   202 (*.for simplify_Integral adapted from 'norm_Rational'.*)
   203 val norm_Rational_noadd_fractions = 
   204    Rule_Set.Sequence {id = "norm_Rational_noadd_fractions", preconds = [], 
   205        rew_ord = ("dummy_ord",Rewrite_Ord.dummy_ord), 
   206        erls = norm_rat_erls, srls = Rule_Set.Empty, calc = [], errpatts = [],
   207        rules = [Rule.Rls_ discard_minus,
   208 		Rule.Rls_ rat_mult_poly,(* removes double fractions like a/b/c    *)
   209 		Rule.Rls_ make_rat_poly_with_parentheses, (*WN0510 also in(#)below*)
   210 		Rule.Rls_ cancel_p_rls, (*FIXME.MG:cancel_p does NOT order sometim*)
   211 		Rule.Rls_ norm_Rational_rls_noadd_fractions,(* the main rls (#)   *)
   212 		Rule.Rls_ discard_parentheses1 (* mult only                       *)
   213 		],
   214        scr = Rule.Empty_Prog
   215        };
   216 
   217 (*.simplify terms before and after Integration such that  
   218    ..a.x^2/2 + b.x^3/3.. is made to ..a/2.x^2 + b/3.x^3.. (and NO
   219    common denominator as done by norm_Rational or make_ratpoly_in.
   220    This is a copy from 'make_ratpoly_in' with respective reduction of rules and
   221    *1* expand the term, ie. distribute * and / over +
   222 .*)
   223 val separate_bdv2 =
   224     Rule_Set.append_rules "separate_bdv2"
   225 	       collect_bdv
   226 	       [\<^rule_thm>\<open>separate_bdv\<close>,
   227 		(*"?a * ?bdv / ?b = ?a / ?b * ?bdv"*)
   228 		\<^rule_thm>\<open>separate_bdv_n\<close>,
   229 		\<^rule_thm>\<open>separate_1_bdv\<close>,
   230 		(*"?bdv / ?b = (1 / ?b) * ?bdv"*)
   231 		\<^rule_thm>\<open>separate_1_bdv_n\<close>(*,
   232 			  (*"?bdv  \<up>  ?n / ?b = 1 / ?b * ?bdv  \<up>  ?n"*)
   233 			  *****\<^rule_thm>\<open>add_divide_distrib\<close>
   234 			  (*"(?x + ?y) / ?z = ?x / ?z + ?y / ?z"*)----------*)
   235 		];
   236 val simplify_Integral = 
   237   Rule_Set.Sequence {id = "simplify_Integral", preconds = []:term list, 
   238        rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
   239       erls = Atools_erls, srls = Rule_Set.Empty,
   240       calc = [],  errpatts = [],
   241       rules = [\<^rule_thm>\<open>distrib_right\<close>,
   242  	       (*"(?z1.0 + ?z2.0) * ?w = ?z1.0 * ?w + ?z2.0 * ?w"*)
   243 	       \<^rule_thm>\<open>add_divide_distrib\<close>,
   244  	       (*"(?x + ?y) / ?z = ?x / ?z + ?y / ?z"*)
   245 	       (*^^^^^ *1* ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^*)
   246 	       Rule.Rls_ norm_Rational_noadd_fractions,
   247 	       Rule.Rls_ order_add_mult_in,
   248 	       Rule.Rls_ discard_parentheses,
   249 	       (*Rule.Rls_ collect_bdv, from make_polynomial_in*)
   250 	       Rule.Rls_ separate_bdv2,
   251 	       \<^rule_eval>\<open>divide\<close> (Prog_Expr.eval_cancel "#divide_e")
   252 	       ],
   253       scr = Rule.Empty_Prog};      
   254 
   255 
   256 (*simplify terms before and after Integration such that  
   257    ..a.x^2/2 + b.x^3/3.. is made to ..a/2.x^2 + b/3.x^3.. (and NO
   258    common denominator as done by norm_Rational or make_ratpoly_in.
   259    This is a copy from 'make_polynomial_in' with insertions from 
   260    'make_ratpoly_in' 
   261 THIS IS KEPT FOR COMPARISON ............................................   
   262 * val simplify_Integral = prep_rls'(
   263 *   Rule_Set.Sequence {id = "", preconds = []:term list, 
   264 *        rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
   265 *       erls = Atools_erls, srls = Rule_Set.Empty,
   266 *       calc = [], (*asm_thm = [],*)
   267 *       rules = [Rule.Rls_ expand_poly,
   268 * 	       Rule.Rls_ order_add_mult_in,
   269 * 	       Rule.Rls_ simplify_power,
   270 * 	       Rule.Rls_ collect_numerals,
   271 * 	       Rule.Rls_ reduce_012,
   272 * 	       \<^rule_thm>\<open>realpow_oneI\<close>,
   273 * 	       Rule.Rls_ discard_parentheses,
   274 * 	       Rule.Rls_ collect_bdv,
   275 * 	       (*below inserted from 'make_ratpoly_in'*)
   276 * 	       Rule.Rls_ (Rule_Set.append_rules "separate_bdv"
   277 * 			 collect_bdv
   278 * 			 [\<^rule_thm>\<open>separate_bdv\<close>,
   279 * 			  (*"?a * ?bdv / ?b = ?a / ?b * ?bdv"*)
   280 * 			  \<^rule_thm>\<open>separate_bdv_n\<close>,
   281 * 			  \<^rule_thm>\<open>separate_1_bdv\<close>,
   282 * 			  (*"?bdv / ?b = (1 / ?b) * ?bdv"*)
   283 * 			  \<^rule_thm>\<open>separate_1_bdv_n\<close>(*,
   284 * 			  (*"?bdv  \<up>  ?n / ?b = 1 / ?b * ?bdv  \<up>  ?n"*)
   285 * 			  \<^rule_thm>\<open>add_divide_distrib\<close>
   286 * 			   (*"(?x + ?y) / ?z = ?x / ?z + ?y / ?z"*)*)
   287 * 			  ]),
   288 * 	       \<^rule_eval>\<open>divide\<close> (eval_cancel "#divide_e")
   289 * 	       ],
   290 *       scr = Rule.Empty_Prog
   291 *       }); 
   292 .......................................................................*)
   293 
   294 val integration = 
   295     Rule_Set.Sequence {id="integration", preconds = [], 
   296 	 rew_ord = ("termlessI",termlessI), 
   297 	 erls = Rule_Def.Repeat {id="conditions_in_integration", 
   298 		     preconds = [], 
   299 		     rew_ord = ("termlessI",termlessI), 
   300 		     erls = Rule_Set.Empty, 
   301 		     srls = Rule_Set.Empty, calc = [], errpatts = [],
   302 		     rules = [],
   303 		     scr = Rule.Empty_Prog}, 
   304 	 srls = Rule_Set.Empty, calc = [], errpatts = [],
   305 	 rules = [ Rule.Rls_ integration_rules,
   306 		   Rule.Rls_ add_new_c,
   307 		   Rule.Rls_ simplify_Integral
   308 		   ],
   309 	 scr = Rule.Empty_Prog};
   310 
   311 val prep_rls' = Auto_Prog.prep_rls @{theory};
   312 \<close>
   313 rule_set_knowledge
   314   integration_rules = \<open>prep_rls' integration_rules\<close> and
   315   add_new_c = \<open>prep_rls' add_new_c\<close> and
   316   simplify_Integral = \<open>prep_rls' simplify_Integral\<close> and
   317   integration = \<open>prep_rls' integration\<close> and
   318   separate_bdv2 = \<open>prep_rls' separate_bdv2\<close> and
   319   norm_Rational_noadd_fractions = \<open>prep_rls' norm_Rational_noadd_fractions\<close> and
   320   norm_Rational_rls_noadd_fractions = \<open>prep_rls' norm_Rational_rls_noadd_fractions\<close>
   321 
   322 (** problems **)
   323 
   324 problem pbl_fun_integ : "integrate/function" =
   325   \<open>Rule_Set.append_rules "empty" Rule_Set.empty [(*for preds in where_*)]\<close>
   326   Method: "diff/integration"
   327   CAS: "Integrate (f_f, v_v)"
   328   Given: "functionTerm f_f" "integrateBy v_v"
   329   Find: "antiDerivative F_F"
   330 
   331 problem pbl_fun_integ_nam : "named/integrate/function" =
   332   (*here "named" is used differently from Differentiation"*)
   333   \<open>Rule_Set.append_rules "empty" Rule_Set.empty [(*for preds in where_*)]\<close>
   334   Method: "diff/integration/named"
   335   CAS: "Integrate (f_f, v_v)"
   336   Given: "functionTerm f_f" "integrateBy v_v"
   337   Find: "antiDerivativeName F_F"
   338 
   339 (** methods **)
   340 
   341 partial_function (tailrec) integrate :: "real \<Rightarrow> real \<Rightarrow> real"
   342   where
   343 "integrate f_f v_v = (
   344   let
   345     t_t = Take (Integral f_f D v_v)
   346   in
   347     (Rewrite_Set_Inst [(''bdv'', v_v)] ''integration'') t_t)"
   348 
   349 method met_diffint : "diff/integration" =
   350   \<open>{rew_ord'="tless_true", rls'=Atools_erls, calc = [], srls = Rule_Set.empty, prls=Rule_Set.empty,
   351 	  crls = Atools_erls, errpats = [], nrls = Rule_Set.empty}\<close>
   352   Program: integrate.simps
   353   Given: "functionTerm f_f" "integrateBy v_v"
   354   Find: "antiDerivative F_F"
   355 
   356 partial_function (tailrec) intergrate_named :: "real \<Rightarrow> real \<Rightarrow> (real \<Rightarrow> real) \<Rightarrow> bool"
   357   where
   358 "intergrate_named f_f v_v F_F =(
   359   let
   360     t_t = Take (F_F v_v = Integral f_f D v_v)
   361   in (
   362     (Try (Rewrite_Set_Inst [(''bdv'', v_v)] ''simplify_Integral'')) #>
   363     (Rewrite_Set_Inst [(''bdv'', v_v)] ''integration'')
   364     ) t_t)"
   365 
   366 method met_diffint_named : "diff/integration/named" =
   367   \<open>{rew_ord'="tless_true", rls'=Atools_erls, calc = [], srls = Rule_Set.empty, prls=Rule_Set.empty,
   368     crls = Atools_erls, errpats = [], nrls = Rule_Set.empty}\<close>
   369   Program: intergrate_named.simps
   370   Given: "functionTerm f_f" "integrateBy v_v"
   371   Find: "antiDerivativeName F_F"
   372 
   373 ML \<open>
   374 \<close> ML \<open>
   375 \<close>
   376 
   377 end