src/Tools/isac/Knowledge/Integrate.thy
author Walther Neuper <wneuper@ist.tugraz.at>
Sat, 10 Mar 2018 17:20:15 +0100
changeset 59399 ca7bdb7da417
parent 59390 f6374c995ac5
child 59400 ef7885190ee8
permissions -rw-r--r--
TermC: clean tests, partially
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@37996
    16
  integrateBy         :: "real => una"
neuper@37996
    17
  antiDerivative      :: "real => una"
neuper@37996
    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@52148
    29
axiomatization where
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@52148
    34
  integral_const:    "Not (bdv occurs_in u) ==> Integral u D bdv = u * bdv" and
neuper@52148
    35
  integral_var:      "Integral bdv D bdv = bdv ^^^ 2 / 2" and
neuper@37906
    36
neuper@37983
    37
  integral_add:      "Integral (u + v) D bdv =  
neuper@52148
    38
		     (Integral u D bdv) + (Integral v D bdv)" and
neuper@37983
    39
  integral_mult:     "[| Not (bdv occurs_in u); bdv occurs_in v |] ==>  
neuper@52148
    40
		     Integral (u * v) D bdv = u * (Integral v D bdv)" and
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@40836
    59
	    case (Symbol.explode o id_of) var of
neuper@37954
    60
		"c"::[] => true
wneuper@59390
    61
	      |	"c"::"_"::is => (case (TermC.int_of_str_opt o implode) is of
neuper@37954
    62
				     SOME _ => true
neuper@37954
    63
				   | NONE => false)
neuper@37954
    64
              | _ => false;
neuper@40836
    65
	fun get_coeff c = case (Symbol.explode o id_of) c of
wneuper@59390
    66
	      		      "c"::"_"::is => (the o TermC.int_of_str_opt o implode) is
neuper@37954
    67
			    | _ => 0;
wneuper@59389
    68
        val cs = filter selc (TermC.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@41922
    92
		     Const ("HOL.eq", T) $ lh $ rh => 
wneuper@59389
    93
		     Const ("HOL.eq", T) $ lh $ TermC.mk_add rh (new_c rh)
wneuper@59389
    94
		   | p => TermC.mk_add p (new_c p)
neuper@37954
    95
    in SOME ((term2str p) ^ " = " ^ term2str p',
wneuper@59390
    96
	  HOLogic.Trueprop $ (TermC.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)) _ =
wneuper@59389
   104
    if TermC.is_f_x arg
neuper@37954
   105
    then SOME ((term2str p) ^ " = True",
wneuper@59390
   106
	       HOLogic.Trueprop $ (TermC.mk_equality (p, @{term True})))
neuper@37954
   107
    else SOME ((term2str p) ^ " = False",
wneuper@59390
   108
	       HOLogic.Trueprop $ (TermC.mk_equality (p, @{term False})))
neuper@37954
   109
  | eval_is_f_x _ _ _ _ = NONE;
neuper@37996
   110
*}
s1210629013@52145
   111
setup {* KEStore_Elems.add_calcs
s1210629013@52145
   112
  [("add_new_c", ("Integrate.add'_new'_c", eval_add_new_c "add_new_c_")),
s1210629013@52145
   113
    ("is_f_x", ("Integrate.is'_f'_x", eval_is_f_x "is_f_idextifier_"))] *}
neuper@37996
   114
ML {*
neuper@37954
   115
(** rulesets **)
neuper@37954
   116
neuper@37954
   117
(*.rulesets for integration.*)
neuper@37954
   118
val integration_rules = 
neuper@37954
   119
    Rls {id="integration_rules", preconds = [], 
neuper@37954
   120
	 rew_ord = ("termlessI",termlessI), 
neuper@37954
   121
	 erls = Rls {id="conditions_in_integration_rules", 
neuper@37954
   122
		     preconds = [], 
neuper@37954
   123
		     rew_ord = ("termlessI",termlessI), 
neuper@37954
   124
		     erls = Erls, 
neuper@42451
   125
		     srls = Erls, calc = [], errpatts = [],
neuper@37954
   126
		     rules = [(*for rewriting conditions in Thm's*)
neuper@37954
   127
			      Calc ("Atools.occurs'_in", 
neuper@37954
   128
				    eval_occurs_in "#occurs_in_"),
wneuper@59389
   129
			      Thm ("not_true", @{thm not_true}),
neuper@37969
   130
			      Thm ("not_false",@{thm not_false})
neuper@37954
   131
			      ],
neuper@37954
   132
		     scr = EmptyScr}, 
neuper@42451
   133
	 srls = Erls, calc = [], errpatts = [],
neuper@37954
   134
	 rules = [
wneuper@59389
   135
		  Thm ("integral_const", @{thm integral_const}),
wneuper@59389
   136
		  Thm ("integral_var", @{thm integral_var}),
wneuper@59389
   137
		  Thm ("integral_add", @{thm integral_add}),
wneuper@59389
   138
		  Thm ("integral_mult", @{thm integral_mult}),
wneuper@59399
   139
		  Thm ("integral_pow", TermC.num_str @{thm integral_pow}),
neuper@38014
   140
		  Calc ("Groups.plus_class.plus", eval_binop "#add_")(*for n+1*)
neuper@37954
   141
		  ],
neuper@37954
   142
	 scr = EmptyScr};
neuper@37996
   143
*}
neuper@37996
   144
ML {*
neuper@37954
   145
val add_new_c = 
neuper@37954
   146
    Seq {id="add_new_c", preconds = [], 
neuper@37954
   147
	 rew_ord = ("termlessI",termlessI), 
neuper@37954
   148
	 erls = Rls {id="conditions_in_add_new_c", 
neuper@37954
   149
		     preconds = [], 
neuper@37954
   150
		     rew_ord = ("termlessI",termlessI), 
neuper@37954
   151
		     erls = Erls, 
neuper@42451
   152
		     srls = Erls, calc = [], errpatts = [],
neuper@37954
   153
		     rules = [Calc ("Tools.matches", eval_matches""),
neuper@37954
   154
			      Calc ("Integrate.is'_f'_x", 
neuper@37954
   155
				    eval_is_f_x "is_f_x_"),
wneuper@59389
   156
			      Thm ("not_true", @{thm not_true}),
wneuper@59389
   157
			      Thm ("not_false", @{thm not_false})
neuper@37954
   158
			      ],
neuper@37954
   159
		     scr = EmptyScr}, 
neuper@42451
   160
	 srls = Erls, calc = [], errpatts = [],
wneuper@59389
   161
	 rules = [ (*Thm ("call_for_new_c",  @{thm call_for_new_c}),*)
neuper@37954
   162
		   Cal1 ("Integrate.add'_new'_c", eval_add_new_c "new_c_")
neuper@37954
   163
		   ],
neuper@37954
   164
	 scr = EmptyScr};
neuper@37996
   165
*}
neuper@37996
   166
ML {*
neuper@37954
   167
neuper@37954
   168
(*.rulesets for simplifying Integrals.*)
neuper@37954
   169
neuper@37954
   170
(*.for simplify_Integral adapted from 'norm_Rational_rls'.*)
neuper@37954
   171
val norm_Rational_rls_noadd_fractions = 
neuper@37954
   172
Rls {id = "norm_Rational_rls_noadd_fractions", preconds = [], 
neuper@37954
   173
     rew_ord = ("dummy_ord",dummy_ord), 
neuper@42451
   174
     erls = norm_rat_erls, srls = Erls, calc = [], errpatts = [],
neuper@52105
   175
     rules = [(*Rls_ add_fractions_p_rls,!!!*)
neuper@37954
   176
	      Rls_ (*rat_mult_div_pow original corrected WN051028*)
neuper@37954
   177
		  (Rls {id = "rat_mult_div_pow", preconds = [], 
neuper@37954
   178
		       rew_ord = ("dummy_ord",dummy_ord), 
neuper@37954
   179
		       erls = (*FIXME.WN051028 e_rls,*)
neuper@37954
   180
		       append_rls "e_rls-is_polyexp" e_rls
neuper@37954
   181
				  [Calc ("Poly.is'_polyexp", 
neuper@37954
   182
					 eval_is_polyexp "")],
neuper@42451
   183
				  srls = Erls, calc = [], errpatts = [],
wneuper@59389
   184
				  rules = [Thm ("rat_mult", @{thm rat_mult}),
neuper@37954
   185
	       (*"?a / ?b * (?c / ?d) = ?a * ?c / (?b * ?d)"*)
wneuper@59389
   186
	       Thm ("rat_mult_poly_l", @{thm rat_mult_poly_l}),
neuper@37954
   187
	       (*"?c is_polyexp ==> ?c * (?a / ?b) = ?c * ?a / ?b"*)
wneuper@59389
   188
	       Thm ("rat_mult_poly_r", @{thm rat_mult_poly_r}),
neuper@37954
   189
	       (*"?c is_polyexp ==> ?a / ?b * ?c = ?a * ?c / ?b"*)
neuper@37954
   190
neuper@37979
   191
	       Thm ("real_divide_divide1_mg",
wneuper@59389
   192
                      @{thm real_divide_divide1_mg}),
neuper@37954
   193
	       (*"y ~= 0 ==> (u / v) / (y / z) = (u * z) / (y * v)"*)
neuper@37979
   194
	       Thm ("divide_divide_eq_right", 
wneuper@59389
   195
                      @{thm divide_divide_eq_right}),
neuper@37954
   196
	       (*"?x / (?y / ?z) = ?x * ?z / ?y"*)
neuper@37979
   197
	       Thm ("divide_divide_eq_left",
wneuper@59389
   198
                      @{thm divide_divide_eq_left}),
neuper@37954
   199
	       (*"?x / ?y / ?z = ?x / (?y * ?z)"*)
wneuper@59360
   200
	       Calc ("Rings.divide_class.divide"  ,eval_cancel "#divide_e"),
neuper@37954
   201
	      
wneuper@59389
   202
	       Thm ("rat_power",  @{thm rat_power})
neuper@37954
   203
		(*"(?a / ?b) ^^^ ?n = ?a ^^^ ?n / ?b ^^^ ?n"*)
neuper@37954
   204
	       ],
wneuper@59389
   205
      scr = Prog ((Thm.term_of o the o (TermC.parse thy)) "empty_script")
neuper@37954
   206
      }),
neuper@37954
   207
		Rls_ make_rat_poly_with_parentheses,
neuper@37954
   208
		Rls_ cancel_p_rls,(*FIXME:cancel_p does NOT order sometimes*)
neuper@37954
   209
		Rls_ rat_reduce_1
neuper@37954
   210
		],
wneuper@59389
   211
       scr = Prog ((Thm.term_of o the o (TermC.parse thy)) "empty_script")
neuper@37954
   212
       }:rls;
neuper@37954
   213
neuper@37954
   214
(*.for simplify_Integral adapted from 'norm_Rational'.*)
neuper@37954
   215
val norm_Rational_noadd_fractions = 
neuper@37954
   216
   Seq {id = "norm_Rational_noadd_fractions", preconds = [], 
neuper@37954
   217
       rew_ord = ("dummy_ord",dummy_ord), 
neuper@42451
   218
       erls = norm_rat_erls, srls = Erls, calc = [], errpatts = [],
neuper@37980
   219
       rules = [Rls_ discard_minus,
neuper@37954
   220
		Rls_ rat_mult_poly,(* removes double fractions like a/b/c    *)
neuper@37954
   221
		Rls_ make_rat_poly_with_parentheses, (*WN0510 also in(#)below*)
neuper@37954
   222
		Rls_ cancel_p_rls, (*FIXME.MG:cancel_p does NOT order sometim*)
neuper@37954
   223
		Rls_ norm_Rational_rls_noadd_fractions,(* the main rls (#)   *)
neuper@37979
   224
		Rls_ discard_parentheses1 (* mult only                       *)
neuper@37954
   225
		],
wneuper@59389
   226
       scr = Prog ((Thm.term_of o the o (TermC.parse thy)) "empty_script")
neuper@37954
   227
       }:rls;
neuper@37954
   228
neuper@37954
   229
(*.simplify terms before and after Integration such that  
neuper@37954
   230
   ..a.x^2/2 + b.x^3/3.. is made to ..a/2.x^2 + b/3.x^3.. (and NO
neuper@37954
   231
   common denominator as done by norm_Rational or make_ratpoly_in.
neuper@37954
   232
   This is a copy from 'make_ratpoly_in' with respective reduction of rules and
neuper@37954
   233
   *1* expand the term, ie. distribute * and / over +
neuper@37954
   234
.*)
neuper@37954
   235
val separate_bdv2 =
neuper@37954
   236
    append_rls "separate_bdv2"
neuper@37954
   237
	       collect_bdv
wneuper@59389
   238
	       [Thm ("separate_bdv",  @{thm separate_bdv}),
neuper@37954
   239
		(*"?a * ?bdv / ?b = ?a / ?b * ?bdv"*)
wneuper@59389
   240
		Thm ("separate_bdv_n",  @{thm separate_bdv_n}),
wneuper@59389
   241
		Thm ("separate_1_bdv",  @{thm separate_1_bdv}),
neuper@37954
   242
		(*"?bdv / ?b = (1 / ?b) * ?bdv"*)
wneuper@59389
   243
		Thm ("separate_1_bdv_n",  @{thm separate_1_bdv_n})(*,
neuper@37954
   244
			  (*"?bdv ^^^ ?n / ?b = 1 / ?b * ?bdv ^^^ ?n"*)
neuper@37991
   245
			  *****Thm ("add_divide_distrib", 
wneuper@59389
   246
			  ***** @{thm add_divide_distrib})
neuper@37954
   247
			  (*"(?x + ?y) / ?z = ?x / ?z + ?y / ?z"*)----------*)
neuper@37954
   248
		];
neuper@37954
   249
val simplify_Integral = 
neuper@37954
   250
  Seq {id = "simplify_Integral", preconds = []:term list, 
neuper@37954
   251
       rew_ord = ("dummy_ord", dummy_ord),
neuper@37954
   252
      erls = Atools_erls, srls = Erls,
neuper@42451
   253
      calc = [],  errpatts = [],
wneuper@59389
   254
      rules = [Thm ("distrib_right", @{thm distrib_right}),
neuper@37954
   255
 	       (*"(?z1.0 + ?z2.0) * ?w = ?z1.0 * ?w + ?z2.0 * ?w"*)
wneuper@59389
   256
	       Thm ("add_divide_distrib", @{thm add_divide_distrib}),
neuper@37954
   257
 	       (*"(?x + ?y) / ?z = ?x / ?z + ?y / ?z"*)
neuper@37954
   258
	       (*^^^^^ *1* ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^*)
neuper@37954
   259
	       Rls_ norm_Rational_noadd_fractions,
neuper@37954
   260
	       Rls_ order_add_mult_in,
neuper@37954
   261
	       Rls_ discard_parentheses,
neuper@37954
   262
	       (*Rls_ collect_bdv, from make_polynomial_in*)
neuper@37954
   263
	       Rls_ separate_bdv2,
wneuper@59360
   264
	       Calc ("Rings.divide_class.divide"  ,eval_cancel "#divide_e")
neuper@37954
   265
	       ],
neuper@37954
   266
      scr = EmptyScr}:rls;      
neuper@37954
   267
neuper@37954
   268
neuper@37954
   269
(*simplify terms before and after Integration such that  
neuper@37954
   270
   ..a.x^2/2 + b.x^3/3.. is made to ..a/2.x^2 + b/3.x^3.. (and NO
neuper@37954
   271
   common denominator as done by norm_Rational or make_ratpoly_in.
neuper@37954
   272
   This is a copy from 'make_polynomial_in' with insertions from 
neuper@37954
   273
   'make_ratpoly_in' 
neuper@37954
   274
THIS IS KEPT FOR COMPARISON ............................................   
s1210629013@55444
   275
* val simplify_Integral = prep_rls'(
neuper@37954
   276
*   Seq {id = "", preconds = []:term list, 
neuper@37954
   277
*        rew_ord = ("dummy_ord", dummy_ord),
neuper@37954
   278
*       erls = Atools_erls, srls = Erls,
neuper@37954
   279
*       calc = [], (*asm_thm = [],*)
neuper@37954
   280
*       rules = [Rls_ expand_poly,
neuper@37954
   281
* 	       Rls_ order_add_mult_in,
neuper@37954
   282
* 	       Rls_ simplify_power,
neuper@37954
   283
* 	       Rls_ collect_numerals,
neuper@37954
   284
* 	       Rls_ reduce_012,
wneuper@59389
   285
* 	       Thm ("realpow_oneI", @{thm realpow_oneI}),
neuper@37954
   286
* 	       Rls_ discard_parentheses,
neuper@37954
   287
* 	       Rls_ collect_bdv,
neuper@37954
   288
* 	       (*below inserted from 'make_ratpoly_in'*)
neuper@37954
   289
* 	       Rls_ (append_rls "separate_bdv"
neuper@37954
   290
* 			 collect_bdv
wneuper@59389
   291
* 			 [Thm ("separate_bdv",  @{thm separate_bdv}),
neuper@37954
   292
* 			  (*"?a * ?bdv / ?b = ?a / ?b * ?bdv"*)
wneuper@59389
   293
* 			  Thm ("separate_bdv_n",  @{thm separate_bdv_n}),
wneuper@59389
   294
* 			  Thm ("separate_1_bdv",  @{thm separate_1_bdv}),
neuper@37954
   295
* 			  (*"?bdv / ?b = (1 / ?b) * ?bdv"*)
wneuper@59389
   296
* 			  Thm ("separate_1_bdv_n",  @{thm separate_1_bdv_n})(*,
neuper@37954
   297
* 			  (*"?bdv ^^^ ?n / ?b = 1 / ?b * ?bdv ^^^ ?n"*)
neuper@37991
   298
* 			  Thm ("add_divide_distrib", 
wneuper@59389
   299
* 				  @{thm add_divide_distrib})
neuper@37954
   300
* 			   (*"(?x + ?y) / ?z = ?x / ?z + ?y / ?z"*)*)
neuper@37954
   301
* 			  ]),
wneuper@59360
   302
* 	       Calc ("Rings.divide_class.divide"  ,eval_cancel "#divide_e")
neuper@37954
   303
* 	       ],
neuper@37954
   304
*       scr = EmptyScr
neuper@37954
   305
*       }:rls); 
neuper@37954
   306
.......................................................................*)
neuper@37954
   307
neuper@37954
   308
val integration = 
neuper@37954
   309
    Seq {id="integration", preconds = [], 
neuper@37954
   310
	 rew_ord = ("termlessI",termlessI), 
neuper@37954
   311
	 erls = Rls {id="conditions_in_integration", 
neuper@37954
   312
		     preconds = [], 
neuper@37954
   313
		     rew_ord = ("termlessI",termlessI), 
neuper@37954
   314
		     erls = Erls, 
neuper@42451
   315
		     srls = Erls, calc = [], errpatts = [],
neuper@37954
   316
		     rules = [],
neuper@37954
   317
		     scr = EmptyScr}, 
neuper@42451
   318
	 srls = Erls, calc = [], errpatts = [],
neuper@37954
   319
	 rules = [ Rls_ integration_rules,
neuper@37954
   320
		   Rls_ add_new_c,
neuper@37954
   321
		   Rls_ simplify_Integral
neuper@37954
   322
		   ],
neuper@37954
   323
	 scr = EmptyScr};
s1210629013@55444
   324
wneuper@59374
   325
val prep_rls' = LTool.prep_rls @{theory};
neuper@37996
   326
*}
neuper@52125
   327
setup {* KEStore_Elems.add_rlss 
s1210629013@55444
   328
  [("integration_rules", (Context.theory_name @{theory}, prep_rls' integration_rules)), 
s1210629013@55444
   329
  ("add_new_c", (Context.theory_name @{theory}, prep_rls' add_new_c)), 
s1210629013@55444
   330
  ("simplify_Integral", (Context.theory_name @{theory}, prep_rls' simplify_Integral)), 
s1210629013@55444
   331
  ("integration", (Context.theory_name @{theory}, prep_rls' integration)), 
s1210629013@55444
   332
  ("separate_bdv2", (Context.theory_name @{theory}, prep_rls' separate_bdv2)),
neuper@52125
   333
neuper@52125
   334
  ("norm_Rational_noadd_fractions", (Context.theory_name @{theory},
s1210629013@55444
   335
    prep_rls' norm_Rational_noadd_fractions)), 
neuper@52125
   336
  ("norm_Rational_rls_noadd_fractions", (Context.theory_name @{theory},
s1210629013@55444
   337
    prep_rls' norm_Rational_rls_noadd_fractions))] *}
neuper@37954
   338
neuper@37954
   339
(** problems **)
s1210629013@55359
   340
setup {* KEStore_Elems.add_pbts
wneuper@59269
   341
  [(Specify.prep_pbt thy "pbl_fun_integ" [] e_pblID
s1210629013@55339
   342
      (["integrate","function"],
s1210629013@55339
   343
        [("#Given" ,["functionTerm f_f", "integrateBy v_v"]),
s1210629013@55339
   344
          ("#Find"  ,["antiDerivative F_F"])],
s1210629013@55339
   345
        append_rls "e_rls" e_rls [(*for preds in where_*)], 
s1210629013@55339
   346
        SOME "Integrate (f_f, v_v)", 
s1210629013@55339
   347
        [["diff","integration"]])),
s1210629013@55339
   348
    (*here "named" is used differently from Differentiation"*)
wneuper@59269
   349
    (Specify.prep_pbt thy "pbl_fun_integ_nam" [] e_pblID
s1210629013@55339
   350
      (["named","integrate","function"],
s1210629013@55339
   351
        [("#Given" ,["functionTerm f_f", "integrateBy v_v"]),
s1210629013@55339
   352
          ("#Find"  ,["antiDerivativeName F_F"])],
s1210629013@55339
   353
        append_rls "e_rls" e_rls [(*for preds in where_*)], 
s1210629013@55339
   354
        SOME "Integrate (f_f, v_v)", 
s1210629013@55339
   355
        [["diff","integration","named"]]))] *}
s1210629013@55380
   356
neuper@37954
   357
(** methods **)
s1210629013@55373
   358
setup {* KEStore_Elems.add_mets
wneuper@59269
   359
  [Specify.prep_met thy "met_diffint" [] e_metID
s1210629013@55373
   360
	    (["diff","integration"],
s1210629013@55373
   361
	      [("#Given" ,["functionTerm f_f", "integrateBy v_v"]), ("#Find"  ,["antiDerivative F_F"])],
s1210629013@55373
   362
	      {rew_ord'="tless_true", rls'=Atools_erls, calc = [], srls = e_rls, prls=e_rls,
s1210629013@55373
   363
	        crls = Atools_erls, errpats = [], nrls = e_rls},
s1210629013@55373
   364
	      "Script IntegrationScript (f_f::real) (v_v::real) =                " ^
s1210629013@55373
   365
          "  (let t_t = Take (Integral f_f D v_v)                             " ^
s1210629013@55373
   366
          "   in (Rewrite_Set_Inst [(bdv,v_v)] integration False) (t_t::real))"),
wneuper@59269
   367
    Specify.prep_met thy "met_diffint_named" [] e_metID
s1210629013@55373
   368
	    (["diff","integration","named"],
s1210629013@55373
   369
	      [("#Given" ,["functionTerm f_f", "integrateBy v_v"]),
s1210629013@55373
   370
	        ("#Find"  ,["antiDerivativeName F_F"])],
s1210629013@55373
   371
	      {rew_ord'="tless_true", rls'=Atools_erls, calc = [], srls = e_rls, prls=e_rls,
s1210629013@55373
   372
          crls = Atools_erls, errpats = [], nrls = e_rls},
s1210629013@55373
   373
        "Script NamedIntegrationScript (f_f::real) (v_v::real) (F_F::real=>real) = " ^
s1210629013@55373
   374
          "  (let t_t = Take (F_F v_v = Integral f_f D v_v)                            " ^
s1210629013@55373
   375
          "   in ((Try (Rewrite_Set_Inst [(bdv,v_v)] simplify_Integral False)) @@  " ^
s1210629013@55373
   376
          "       (Rewrite_Set_Inst [(bdv,v_v)] integration False)) t_t)            ")]
s1210629013@55373
   377
*}
neuper@37954
   378
neuper@37906
   379
end