src/Tools/isac/Knowledge/Partial_Fractions.thy
author Walther Neuper <neuper@ist.tugraz.at>
Sat, 05 May 2012 14:28:20 +0200
changeset 42415 1df1c4c047cd
parent 42414 016fc77bdcd1
child 42417 fd6ddf0f843a
permissions -rwxr-xr-x
tuned
neuper@42376
     1
(* Partial_Fractions 
jan@42344
     2
   author: Jan Rocnik, isac team
jan@42344
     3
   Copyright (c) isac team 2011
jan@42344
     4
   Use is subject to license terms.
jan@42344
     5
jan@42344
     6
1234567890123456789012345678901234567890123456789012345678901234567890123456789012345678901234567890
jan@42344
     7
        10        20        30        40        50        60        70        80         90      100
jan@42344
     8
*)
neuper@42376
     9
header {* Partial Fraction Decomposition *}
jan@42344
    10
neuper@42289
    11
neuper@42376
    12
theory Partial_Fractions imports RootRatEq begin
jan@42353
    13
jan@42353
    14
ML{*
neuper@42376
    15
(*
jan@42353
    16
signature PARTFRAC =
jan@42353
    17
sig
jan@42353
    18
  val ansatz_rls : rls
jan@42353
    19
  val ansatz_rls_ : theory -> term -> (term * term list) option
jan@42353
    20
end
neuper@42376
    21
*)
jan@42353
    22
*}
jan@42353
    23
neuper@42376
    24
subsection {* eval_ functions *}
jan@42344
    25
consts
neuper@42359
    26
  factors_from_solution :: "bool list => real"
neuper@42359
    27
  drop_questionmarks    :: "'a => 'a"
jan@42344
    28
neuper@42376
    29
text {* these might be used for variants of fac_from_sol *}
neuper@42289
    30
ML {*
jan@42344
    31
fun mk_minus_1 T = Free("-1", T); (*TODO DELETE WITH numbers_to_string*)
jan@42344
    32
fun flip_sign t = (*TODO improve for use in factors_from_solution: -(-1) etc*)
jan@42344
    33
  let val minus_1 = t |> type_of |> mk_minus_1
jan@42344
    34
  in HOLogic.mk_binop "Groups.times_class.times" (minus_1, t) end;
neuper@42376
    35
*}
neuper@42376
    36
neuper@42376
    37
text {* from solutions (e.g. [z = 1, z = -2]) make linear factors (e.g. (z - 1)*(z - -2)) *}
neuper@42376
    38
ML {*
jan@42344
    39
fun fac_from_sol s =
jan@42344
    40
  let val (lhs, rhs) = HOLogic.dest_eq s
jan@42367
    41
  in HOLogic.mk_binop "Groups.minus_class.minus" (lhs, rhs) end;
jan@42344
    42
jan@42344
    43
fun mk_prod prod [] =
jan@42344
    44
      if prod = e_term then error "mk_prod called with []" else prod
jan@42344
    45
  | mk_prod prod (t :: []) =
jan@42344
    46
      if prod = e_term then t else HOLogic.mk_binop "Groups.times_class.times" (prod, t)
jan@42344
    47
  | mk_prod prod (t1 :: t2 :: ts) =
jan@42344
    48
        if prod = e_term 
jan@42344
    49
        then 
jan@42344
    50
           let val p = HOLogic.mk_binop "Groups.times_class.times" (t1, t2)
jan@42344
    51
           in mk_prod p ts end 
jan@42344
    52
        else 
jan@42344
    53
           let val p = HOLogic.mk_binop "Groups.times_class.times" (prod, t1)
jan@42344
    54
           in mk_prod p (t2 :: ts) end 
jan@42344
    55
jan@42344
    56
fun factors_from_solution sol = 
jan@42344
    57
  let val ts = HOLogic.dest_list sol
jan@42344
    58
  in mk_prod e_term (map fac_from_sol ts) end;
jan@42344
    59
neuper@42376
    60
(*("factors_from_solution", ("Partial_Fractions.factors_from_solution", 
neuper@42376
    61
     eval_factors_from_solution ""))*)
jan@42352
    62
fun eval_factors_from_solution (thmid:string) _
jan@42352
    63
     (t as Const ("Partial_Fractions.factors_from_solution", _) $ sol) thy =
jan@42352
    64
       ((let val prod = factors_from_solution sol
jan@42352
    65
         in SOME (mk_thmid thmid ""
jan@42352
    66
           (Print_Mode.setmp [] (Syntax.string_of_term (thy2ctxt thy)) prod) "",
jan@42352
    67
               Trueprop $ (mk_equality (t, prod)))
jan@42352
    68
         end)
jan@42352
    69
       handle _ => NONE)
jan@42352
    70
 | eval_factors_from_solution _ _ _ _ = NONE;
jan@42344
    71
*}
jan@42344
    72
neuper@42376
    73
text {* 'ansatz' introduces '?Vars' (questionable design); drop these again *}
neuper@42359
    74
ML {*
neuper@42359
    75
(*("drop_questionmarks", ("Partial_Fractions.drop_questionmarks", eval_drop_questionmarks ""))*)
neuper@42359
    76
fun eval_drop_questionmarks (thmid:string) _
neuper@42359
    77
     (t as Const ("Partial_Fractions.drop_questionmarks", _) $ tm) thy =
neuper@42359
    78
        if contains_Var tm
neuper@42359
    79
        then
neuper@42359
    80
          let
neuper@42359
    81
            val tm' = var2free tm
neuper@42359
    82
            in SOME (mk_thmid thmid ""
neuper@42359
    83
              (Print_Mode.setmp [] (Syntax.string_of_term (thy2ctxt thy)) tm') "",
neuper@42359
    84
                Trueprop $ (mk_equality (t, tm')))
neuper@42359
    85
            end
neuper@42359
    86
        else NONE
neuper@42359
    87
  | eval_drop_questionmarks _ _ _ _ = NONE;
neuper@42376
    88
*}
neuper@42359
    89
neuper@42376
    90
text {* store eval_ functions for calls from Scripts *}
neuper@42376
    91
ML {*
neuper@42359
    92
calclist':= overwritel (!calclist',
neuper@42359
    93
  [("drop_questionmarks", ("Partial_Fractions.drop'_questionmarks", eval_drop_questionmarks ""))
neuper@42359
    94
  ]);
neuper@42359
    95
*}
neuper@42359
    96
neuper@42376
    97
subsection {* 'ansatz' for partial fractions *}
jan@42353
    98
axiomatization where
jan@42358
    99
  ansatz_2nd_order: "n / (a*b) = A/a + B/b" and
neuper@42376
   100
  ansatz_3rd_order: "n / (a*b*c) = A/a + B/b + C/c" and
neuper@42376
   101
  ansatz_4th_order: "n / (a*b*c*d) = A/a + B/b + C/c + D/d" and
neuper@42386
   102
  (*version 1*)
neuper@42376
   103
  equival_trans_2nd_order: "(n/(a*b) = A/a + B/b) = (n = A*b + B*a)" and
neuper@42376
   104
  equival_trans_3rd_order: "(n/(a*b*c) = A/a + B/b + C/c) = (n = A*b*c + B*a*c + C*a*b)" and
neuper@42376
   105
  equival_trans_4th_order: "(n/(a*b*c*d) = A/a + B/b + C/c + D/d) = 
neuper@42386
   106
    (n = A*b*c*d + B*a*c*d + C*a*b*d + D*a*b*c)" and
neuper@42386
   107
  (*version 2: not yet used, see partial_fractions.sml*)
neuper@42387
   108
  multiply_2nd_order: "(n/x = A/a + B/b) = (a*b*n/x = A*b + B*a)" and
neuper@42387
   109
  multiply_3rd_order: "(n/x = A/a + B/b + C/c) = (a*b*c*n/x = A*b*c + B*a*c + C*a*b)" and
neuper@42387
   110
  multiply_4th_order: 
neuper@42387
   111
    "(n/x = A/a + B/b + C/c + D/d) = (a*b*c*d*n/x = A*b*c*d + B*a*c*d + C*a*b*d + D*a*b*c)"
neuper@42387
   112
neuper@42387
   113
text {* Probably the optimal formalization woudl be ...
neuper@42387
   114
neuper@42386
   115
  multiply_2nd_order: "x = a*b ==> (n/x = A/a + B/b) = (a*b*n/x = A*b + B*a)" and
neuper@42386
   116
  multiply_3rd_order: "x = a*b*c ==>
neuper@42386
   117
    (n/x = A/a + B/b + C/c) = (a*b*c*n/x = A*b*c + B*a*c + C*a*b)" and
neuper@42386
   118
  multiply_4th_order: "x = a*b*c*d ==>
neuper@42386
   119
    (n/x = A/a + B/b + C/c + D/d) = (a*b*c*d*n/x = A*b*c*d + B*a*c*d + C*a*b*d + D*a*b*c)"
jan@42353
   120
neuper@42387
   121
... because it would allow to start the ansatz as follows
neuper@42387
   122
(1) 3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z))) = 3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z)))
neuper@42387
   123
(2) 3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z))) = AA / (z - 1 / 2) + BB / (z - -1 / 4)
neuper@42387
   124
(3) (z - 1 / 2) * (z - -1 / 4) * 3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z))) = 
neuper@42387
   125
    (z - 1 / 2) * (z - -1 / 4) * AA / (z - 1 / 2) + BB / (z - -1 / 4)
neuper@42387
   126
(4) 3 = A * (z - -1 / 4) + B * (z - 1 / 2)
neuper@42387
   127
neuper@42387
   128
... (1==>2) ansatz
neuper@42387
   129
    (2==>3) multiply_*
neuper@42387
   130
    (3==>4) norm_Rational
neuper@42387
   131
TODOs for this version ar in partial_fractions.sml "--- progr.vers.2: "
neuper@42387
   132
*}
neuper@42387
   133
jan@42353
   134
ML {*
jan@42353
   135
val ansatz_rls = prep_rls(
jan@42353
   136
  Rls {id = "ansatz_rls", preconds = [], rew_ord = ("dummy_ord",dummy_ord), 
jan@42353
   137
	  erls = Erls, srls = Erls, calc = [],
jan@42353
   138
	  rules = 
neuper@42376
   139
	   [Thm ("ansatz_2nd_order",num_str @{thm ansatz_2nd_order}),
neuper@42376
   140
	    Thm ("ansatz_3rd_order",num_str @{thm ansatz_3rd_order})
jan@42353
   141
	   ], 
jan@42353
   142
	 scr = EmptyScr}:rls);
jan@42353
   143
jan@42354
   144
val equival_trans = prep_rls(
jan@42354
   145
  Rls {id = "equival_trans", preconds = [], rew_ord = ("dummy_ord",dummy_ord), 
jan@42354
   146
	  erls = Erls, srls = Erls, calc = [],
jan@42354
   147
	  rules = 
neuper@42376
   148
	   [Thm ("equival_trans_2nd_order",num_str @{thm equival_trans_2nd_order}),
neuper@42376
   149
	    Thm ("equival_trans_3rd_order",num_str @{thm equival_trans_3rd_order})
jan@42354
   150
	   ], 
jan@42354
   151
	 scr = EmptyScr}:rls);
neuper@42386
   152
neuper@42386
   153
val multiply_ansatz = prep_rls(
neuper@42386
   154
  Rls {id = "multiply_ansatz", preconds = [], rew_ord = ("dummy_ord",dummy_ord), 
neuper@42386
   155
	  erls = Erls,
neuper@42386
   156
	  srls = Erls, calc = [],
neuper@42386
   157
	  rules = 
neuper@42386
   158
	   [Thm ("multiply_2nd_order",num_str @{thm multiply_2nd_order})
neuper@42386
   159
	   ], 
neuper@42386
   160
	 scr = EmptyScr}:rls);
jan@42354
   161
*}
jan@42354
   162
jan@42354
   163
text {*store the rule set for math engine*}
jan@42354
   164
ML {*
jan@42353
   165
ruleset' := overwritelthy @{theory} (!ruleset',
jan@42354
   166
  [("ansatz_rls", ansatz_rls),
neuper@42386
   167
   ("multiply_ansatz", multiply_ansatz),
jan@42354
   168
   ("equival_trans", equival_trans)
jan@42353
   169
   ]);
jan@42353
   170
*}
jan@42344
   171
neuper@42376
   172
subsection {* Specification *}
jan@42344
   173
neuper@42376
   174
consts
neuper@42376
   175
  decomposedFunction :: "real => una"
neuper@42376
   176
neuper@42289
   177
ML {*
neuper@42386
   178
check_guhs_unique := false; (*WN120307 REMOVE after editing*)
neuper@42376
   179
store_pbt
neuper@42376
   180
 (prep_pbt @{theory} "pbl_simp_rat_partfrac" [] e_pblID
neuper@42376
   181
 (["partial_fraction", "rational", "simplification"],
neuper@42376
   182
  [("#Given" ,["functionTerm t_t", "solveFor v_v"]),
neuper@42376
   183
   (* TODO: call this sub-problem with appropriate functionTerm: 
neuper@42376
   184
      leading coefficient of the denominator is 1: to be checked here! and..
neuper@42376
   185
     ("#Where" ,["((get_numerator t_t) has_degree_in v_v) < 
neuper@42376
   186
                  ((get_denominator t_t) has_degree_in v_v)"]), TODO*)
neuper@42376
   187
   ("#Find"  ,["decomposedFunction p_p"])
neuper@42376
   188
  ],
neuper@42376
   189
  append_rls "e_rls" e_rls [(*for preds in where_ TODO*)], 
neuper@42376
   190
  NONE, 
neuper@42376
   191
  [["simplification","of_rationals","to_partial_fraction"]]));
neuper@42289
   192
*}
jan@42354
   193
neuper@42376
   194
subsection {* Method *}
neuper@42376
   195
consts
neuper@42376
   196
  PartFracScript  :: "[real,real,  real] => real" 
neuper@42376
   197
    ("((Script PartFracScript (_ _ =))// (_))" 9)
jan@42353
   198
neuper@42376
   199
text {* rule set for functions called in the Script *}
neuper@42376
   200
ML {*
neuper@42413
   201
  val srls_partial_fraction = Rls {id="srls_partial_fraction", 
neuper@42376
   202
    preconds = [],
neuper@42376
   203
    rew_ord = ("termlessI",termlessI),
neuper@42376
   204
    erls = append_rls "erls_in_srls_partial_fraction" e_rls
neuper@42376
   205
      [(*for asm in NTH_CONS ...*)
neuper@42376
   206
       Calc ("Orderings.ord_class.less",eval_equ "#less_"),
neuper@42376
   207
       (*2nd NTH_CONS pushes n+-1 into asms*)
neuper@42376
   208
       Calc("Groups.plus_class.plus", eval_binop "#add_")], 
neuper@42376
   209
    srls = Erls, calc = [],
neuper@42376
   210
    rules = [
neuper@42376
   211
       Thm ("NTH_CONS",num_str @{thm NTH_CONS}),
neuper@42376
   212
       Calc("Groups.plus_class.plus", eval_binop "#add_"),
neuper@42376
   213
       Thm ("NTH_NIL",num_str @{thm NTH_NIL}),
neuper@42376
   214
       Calc("Tools.lhs", eval_lhs "eval_lhs_"),
neuper@42376
   215
       Calc("Tools.rhs", eval_rhs"eval_rhs_"),
neuper@42376
   216
       Calc("Atools.argument'_in", eval_argument_in "Atools.argument'_in"),
neuper@42376
   217
       Calc("Rational.get_denominator", eval_get_denominator "#get_denominator"),
neuper@42376
   218
       Calc("Rational.get_numerator", eval_get_numerator "#get_numerator"),
neuper@42376
   219
       Calc("Partial_Fractions.factors_from_solution",
neuper@42376
   220
         eval_factors_from_solution "#factors_from_solution"),
neuper@42376
   221
       Calc("Partial_Fractions.drop_questionmarks", eval_drop_questionmarks "#drop_?")],
neuper@42376
   222
    scr = EmptyScr};
neuper@42376
   223
*}
neuper@42376
   224
ML {*
neuper@42376
   225
eval_drop_questionmarks;
neuper@42376
   226
*}
neuper@42376
   227
ML {*
neuper@42376
   228
val ctxt = ProofContext.init_global @{theory};
neuper@42376
   229
val SOME t = parseNEW ctxt "eqr = drop_questionmarks eqr";
neuper@42376
   230
*}
neuper@42376
   231
ML {*
neuper@42376
   232
parseNEW ctxt "decomposedFunction p_p";
neuper@42376
   233
parseNEW ctxt "decomposedFunction";
neuper@42376
   234
*}
neuper@42376
   235
ML {*
neuper@42376
   236
*}
neuper@42376
   237
ML {* (* current version, error outcommented *)
neuper@42376
   238
store_met
neuper@42376
   239
 (prep_met @{theory} "met_partial_fraction" [] e_metID
neuper@42376
   240
 (["simplification","of_rationals","to_partial_fraction"], 
neuper@42376
   241
  [("#Given" ,["functionTerm t_t", "solveFor v_v"]),
neuper@42376
   242
   (*("#Where" ,["((get_numerator t_t) has_degree_in v_v) < 
neuper@42376
   243
                  ((get_denominator t_t) has_degree_in v_v)"]), TODO*)
neuper@42376
   244
   ("#Find"  ,["decomposedFunction p_p"])],
neuper@42413
   245
   {rew_ord'="tless_true", rls'= e_rls, calc = [], srls = srls_partial_fraction, prls = e_rls,
neuper@42414
   246
    crls = e_rls, nrls = e_rls},                         (*f_f = 3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z)), zzz: z*)
neuper@42414
   247
   "Script PartFracScript (f_f::real) (zzz::real) =   " ^(*([], Frm), Problem (Partial_Fractions, [partial_fraction, rational, simplification])*)
neuper@42414
   248
   "(let f_f = Take f_f;                              " ^(*([1], Frm), 3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z)))*)
neuper@42414
   249
   "  (num_orig::real) = get_numerator f_f;           " ^(*           num_orig = 3*)
neuper@42414
   250
   "  f_f = (Rewrite_Set norm_Rational False) f_f;    " ^(*([1], Res), 24 / (-1 + -2 * z + 8 * z ^^^ 2)*)
neuper@42414
   251
   "  (denom::real) = get_denominator f_f;            " ^(*           denom = -1 + -2 * z + 8 * z ^^^ 2*)
neuper@42414
   252
   "  (equ::bool) = (denom = (0::real));              " ^(*           equ = -1 + -2 * z + 8 * z ^^^ 2 = 0*)
neuper@42415
   253
neuper@42414
   254
   "  (L_L::bool list) = (SubProblem (PolyEq',        " ^(*([2], Pbl), solve (-1 + -2 * z + 8 * z ^^^ 2 = 0, z)*)
neuper@42414
   255
   "    [abcFormula, degree_2, polynomial, univariate, equation], " ^
neuper@42414
   256
   "    [no_met]) [BOOL equ, REAL zzz]);              " ^(*([2], Res), [z = 1 / 2, z = -1 / 4]*)
neuper@42414
   257
   "  (facs::real) = factors_from_solution L_L;       " ^(*           facs: (z - 1 / 2) * (z - -1 / 4)*)
neuper@42414
   258
   "  (eql::real) = Take (num_orig / facs);           " ^(*([3], Frm), 33 / ((z - 1 / 2) * (z - -1 / 4)) *) 
neuper@42414
   259
   "  (eqr::real) = (Try (Rewrite_Set ansatz_rls False)) eql;  " ^(*([3], Res), ?A / (z - 1 / 2) + ?B / (z - -1 / 4)*)
neuper@42414
   260
   "  (eq::bool) = Take (eql = eqr);                  " ^(*([4], Frm), 3 / ((z - 1 / 2) * (z - -1 / 4)) = ?A / (z - 1 / 2) + ?B / (z - -1 / 4)*)
neuper@42414
   261
   "  eq = (Try (Rewrite_Set equival_trans False)) eq;" ^(*([4], Res), 3 = ?A * (z - -1 / 4) + ?B * (z - 1 / 2)*) 
neuper@42414
   262
   "  eq = drop_questionmarks eq;                     " ^(*           eq = 3 = A * (z - -1 / 4) + B * (z - 1 / 2)*)
neuper@42414
   263
   "  (z1::real) = (rhs (NTH 1 L_L));                 " ^(*           z1 = 1 / 2*)
neuper@42414
   264
   "  (z2::real) = (rhs (NTH 2 L_L));                 " ^(*           z2 = -1 / 4*)
neuper@42414
   265
   "  (eq_a::bool) = Take eq;                         " ^(*([5], Frm), 3 = A * (z - -1 / 4) + B * (z - 1 / 2)*)
neuper@42414
   266
   "  eq_a = (Substitute [zzz = z1]) eq;              " ^(*([5], Res), 3 = A * (1 / 2 - -1 / 4) + B * (1 / 2 - 1 / 2)*)
neuper@42414
   267
   "  eq_a = (Rewrite_Set norm_Rational False) eq_a;  " ^(*([6], Res), 3 = 3 * A / 4*)
neuper@42415
   268
neuper@42414
   269
   "  (sol_a::bool list) =                            " ^(*([7], Pbl), solve (3 = 3 * A / 4, A)*)
neuper@42414
   270
   "    (SubProblem (Isac', [univariate,equation], [no_met])   " ^
neuper@42414
   271
   "    [BOOL eq_a, REAL (A::real)]);                 " ^(*([7], Res), [A = 4]*)
neuper@42414
   272
   "  (a::real) = (rhs (NTH 1 sol_a));                " ^(*           a = 4*)
neuper@42414
   273
   "  (eq_b::bool) = Take eq;                         " ^(*([8], Frm), 3 = A * (z - -1 / 4) + B * (z - 1 / 2)*)
neuper@42414
   274
   "  eq_b = (Substitute [zzz = z2]) eq_b;            " ^(*([8], Res), 3 = A * (-1 / 4 - -1 / 4) + B * (-1 / 4 - 1 / 2)*)
neuper@42414
   275
   "  eq_b = (Rewrite_Set norm_Rational False) eq_b;  " ^(*([9], Res), 3 = -3 * B / 4*)
neuper@42414
   276
   "  (sol_b::bool list) =                            " ^(*([10], Pbl), solve (3 = -3 * B / 4, B)*)
neuper@42414
   277
   "    (SubProblem (Isac', [univariate,equation], [no_met])   " ^
neuper@42414
   278
   "    [BOOL eq_b, REAL (B::real)]);                 " ^(*([10], Res), [B = -4]*)
neuper@42414
   279
   "  (b::real) = (rhs (NTH 1 sol_b));                " ^(*           b = -4*)
neuper@42414
   280
   "  eqr = drop_questionmarks eqr;                   " ^(*           eqr = A / (z - 1 / 2) + B / (z - -1 / 4)*)
neuper@42414
   281
   "  (pbz::real) = Take eqr;                         " ^(*([11], Frm), A / (z - 1 / 2) + B / (z - -1 / 4)*)
neuper@42414
   282
   "  pbz = ((Substitute [A = a, B = b]) pbz)         " ^(*([11], Res), 4 / (z - 1 / 2) + -4 / (z - -1 / 4)*)
neuper@42414
   283
   "in pbz)"                                             (*([], Res), 4 / (z - 1 / 2) + -4 / (z - -1 / 4)*)
neuper@42376
   284
));
neuper@42376
   285
*}
neuper@42376
   286
ML {*
neuper@42376
   287
(*
neuper@42376
   288
  val fmz =                                             
neuper@42376
   289
    ["functionTerm (3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z))))", 
neuper@42376
   290
      "solveFor z", "functionTerm p_p"];
neuper@42376
   291
  val (dI',pI',mI') =
neuper@42376
   292
    ("Partial_Fractions", 
neuper@42376
   293
      ["partial_fraction", "rational", "simplification"],
neuper@42376
   294
      ["simplification","of_rationals","to_partial_fraction"]);
neuper@42376
   295
  val (p,_,f,nxt,_,pt) = CalcTreeTEST [(fmz, (dI',pI',mI'))];
neuper@42376
   296
*)
neuper@42376
   297
*}
neuper@42289
   298
jan@42295
   299
neuper@42376
   300
neuper@42376
   301
subsection {**}
neuper@42376
   302
neuper@42289
   303
end