src/Tools/isac/Knowledge/Partial_Fractions.thy
author Walther Neuper <wneuper@ist.tugraz.at>
Tue, 19 Feb 2019 19:35:12 +0100
changeset 59504 546bd1b027cc
parent 59491 516e6cc731ab
child 59505 a1f223658994
permissions -rw-r--r--
[-Test_Isac] funpack: cp program code to partial_function

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