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