src/Tools/isac/Knowledge/Integrate.thy
author Walther Neuper <neuper@ist.tugraz.at>
Wed, 08 Sep 2010 17:20:03 +0200
branchisac-update-Isa09-2
changeset 37994 eb4c556a525b
parent 37991 028442673981
child 37996 eb7d9cbaa3ef
permissions -rw-r--r--
tuned

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