src/Tools/isac/Knowledge/Partial_Fractions.thy
author wneuper <Walther.Neuper@jku.at>
Thu, 04 Aug 2022 12:48:37 +0200
changeset 60509 2e0b7ca391dc
parent 60504 8cc1415b3530
child 60515 03e19793d81e
permissions -rw-r--r--
polish naming in Rewrite_Order
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"
wneuper@59512
    24
  AA                    :: real
wneuper@59512
    25
  BB                    :: real
jan@42344
    26
wneuper@59472
    27
text \<open>these might be used for variants of fac_from_sol\<close>
wneuper@59472
    28
ML \<open>
jan@42344
    29
fun mk_minus_1 T = Free("-1", T); (*TODO DELETE WITH numbers_to_string*)
jan@42344
    30
fun flip_sign t = (*TODO improve for use in factors_from_solution: -(-1) etc*)
jan@42344
    31
  let val minus_1 = t |> type_of |> mk_minus_1
wenzelm@60309
    32
  in HOLogic.mk_binop \<^const_name>\<open>times\<close> (minus_1, t) end;
wneuper@59472
    33
\<close>
neuper@42376
    34
wneuper@59472
    35
text \<open>from solutions (e.g. [z = 1, z = -2]) make linear factors (e.g. (z - 1)*(z - -2))\<close>
wneuper@59472
    36
ML \<open>
jan@42344
    37
fun fac_from_sol s =
jan@42344
    38
  let val (lhs, rhs) = HOLogic.dest_eq s
wenzelm@60309
    39
  in HOLogic.mk_binop \<^const_name>\<open>minus\<close> (lhs, rhs) end;
jan@42344
    40
jan@42344
    41
fun mk_prod prod [] =
walther@59962
    42
      if prod = TermC.empty then raise ERROR "mk_prod called with []" else prod
jan@42344
    43
  | mk_prod prod (t :: []) =
wenzelm@60309
    44
      if prod = TermC.empty then t else HOLogic.mk_binop \<^const_name>\<open>times\<close> (prod, t)
jan@42344
    45
  | mk_prod prod (t1 :: t2 :: ts) =
walther@59861
    46
        if prod = TermC.empty 
jan@42344
    47
        then 
wenzelm@60309
    48
           let val p = HOLogic.mk_binop \<^const_name>\<open>times\<close> (t1, t2)
jan@42344
    49
           in mk_prod p ts end 
jan@42344
    50
        else 
wenzelm@60309
    51
           let val p = HOLogic.mk_binop \<^const_name>\<open>times\<close> (prod, t1)
jan@42344
    52
           in mk_prod p (t2 :: ts) end 
jan@42344
    53
jan@42344
    54
fun factors_from_solution sol = 
jan@42344
    55
  let val ts = HOLogic.dest_list sol
walther@59861
    56
  in mk_prod TermC.empty (map fac_from_sol ts) end;
jan@42344
    57
neuper@42376
    58
(*("factors_from_solution", ("Partial_Fractions.factors_from_solution", 
walther@60264
    59
     eval_factors_from_solution ""))
walther@60264
    60
  this code is limited (max 3 solutions) AND has too little checks *)
jan@42352
    61
fun eval_factors_from_solution (thmid:string) _
Walther@60504
    62
    (t as Const (\<^const_name>\<open>Partial_Fractions.factors_from_solution\<close>, _) $ sol) ctxt =
walther@60264
    63
      let val prod = factors_from_solution sol
Walther@60504
    64
      in SOME (TermC.mk_thmid thmid (UnparseC.term_in_ctxt ctxt prod) "",
walther@60264
    65
           HOLogic.Trueprop $ (TermC.mk_equality (t, prod)))
walther@60264
    66
      end
walther@60264
    67
  | eval_factors_from_solution _ _ _ _ = NONE;
wneuper@59472
    68
\<close>
jan@42344
    69
wneuper@59472
    70
subsection \<open>'ansatz' for partial fractions\<close>
jan@42353
    71
axiomatization where
wneuper@59512
    72
  ansatz_2nd_order: "n / (a*b) = AA/a + BB/b" and
wneuper@59512
    73
  ansatz_3rd_order: "n / (a*b*c) = AA/a + BB/b + C/c" and
wneuper@59512
    74
  ansatz_4th_order: "n / (a*b*c*d) = AA/a + BB/b + C/c + D/d" and
neuper@42386
    75
  (*version 1*)
wneuper@59512
    76
  equival_trans_2nd_order: "(n/(a*b) = AA/a + BB/b) = (n = AA*b + BB*a)" and
wneuper@59512
    77
  equival_trans_3rd_order: "(n/(a*b*c) = AA/a + BB/b + C/c) = (n = AA*b*c + BB*a*c + C*a*b)" and
wneuper@59512
    78
  equival_trans_4th_order: "(n/(a*b*c*d) = AA/a + BB/b + C/c + D/d) = 
wneuper@59512
    79
    (n = AA*b*c*d + BB*a*c*d + C*a*b*d + D*a*b*c)" and
neuper@42386
    80
  (*version 2: not yet used, see partial_fractions.sml*)
wneuper@59512
    81
  multiply_2nd_order: "(n/x = AA/a + BB/b) = (a*b*n/x = AA*b + BB*a)" and
wneuper@59512
    82
  multiply_3rd_order: "(n/x = AA/a + BB/b + C/c) = (a*b*c*n/x = AA*b*c + BB*a*c + C*a*b)" and
neuper@42387
    83
  multiply_4th_order: 
wneuper@59512
    84
    "(n/x = AA/a + BB/b + C/c + D/d) = (a*b*c*d*n/x = AA*b*c*d + BB*a*c*d + C*a*b*d + D*a*b*c)"
neuper@42387
    85
wneuper@59472
    86
text \<open>Probably the optimal formalization woudl be ...
neuper@42387
    87
wneuper@59512
    88
  multiply_2nd_order: "x = a*b ==> (n/x = AA/a + BB/b) = (a*b*n/x = AA*b + BB*a)" and
neuper@42386
    89
  multiply_3rd_order: "x = a*b*c ==>
wneuper@59512
    90
    (n/x = AA/a + BB/b + C/c) = (a*b*c*n/x = AA*b*c + BB*a*c + C*a*b)" and
neuper@42386
    91
  multiply_4th_order: "x = a*b*c*d ==>
wneuper@59512
    92
    (n/x = AA/a + BB/b + C/c + D/d) = (a*b*c*d*n/x = AA*b*c*d + BB*a*c*d + C*a*b*d + D*a*b*c)"
jan@42353
    93
neuper@42387
    94
... because it would allow to start the ansatz as follows
neuper@42387
    95
(1) 3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z))) = 3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z)))
neuper@42387
    96
(2) 3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z))) = AA / (z - 1 / 2) + BB / (z - -1 / 4)
neuper@42387
    97
(3) (z - 1 / 2) * (z - -1 / 4) * 3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z))) = 
neuper@42387
    98
    (z - 1 / 2) * (z - -1 / 4) * AA / (z - 1 / 2) + BB / (z - -1 / 4)
wneuper@59512
    99
(4) 3 = AA * (z - -1 / 4) + BB * (z - 1 / 2)
neuper@42387
   100
neuper@42387
   101
... (1==>2) ansatz
neuper@42387
   102
    (2==>3) multiply_*
neuper@42387
   103
    (3==>4) norm_Rational
wneuper@59550
   104
TODOs for this version are in partial_fractions.sml "--- progr.vers.2: "
wneuper@59472
   105
\<close>
neuper@42387
   106
wneuper@59472
   107
ML \<open>
s1210629013@55444
   108
val ansatz_rls = prep_rls'(
Walther@60509
   109
  Rule_Def.Repeat {id = "ansatz_rls", preconds = [], rew_ord = ("dummy_ord",Rewrite_Ord.function_empty), 
walther@59851
   110
	  erls = Rule_Set.Empty, srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@60358
   111
	  rules = [
walther@60358
   112
      \<^rule_thm>\<open>ansatz_2nd_order\<close>,
walther@60358
   113
	    \<^rule_thm>\<open>ansatz_3rd_order\<close>], 
walther@60358
   114
	  scr = Rule.Empty_Prog});
jan@42353
   115
s1210629013@55444
   116
val equival_trans = prep_rls'(
Walther@60509
   117
  Rule_Def.Repeat {id = "equival_trans", preconds = [], rew_ord = ("dummy_ord",Rewrite_Ord.function_empty), 
walther@59851
   118
	  erls = Rule_Set.Empty, srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@60358
   119
	  rules = [
walther@60358
   120
      \<^rule_thm>\<open>equival_trans_2nd_order\<close>,
walther@60358
   121
	    \<^rule_thm>\<open>equival_trans_3rd_order\<close>], 
walther@60358
   122
	  scr = Rule.Empty_Prog});
neuper@42386
   123
s1210629013@55444
   124
val multiply_ansatz = prep_rls'(
Walther@60509
   125
  Rule_Def.Repeat {id = "multiply_ansatz", preconds = [], rew_ord = ("dummy_ord",Rewrite_Ord.function_empty), 
walther@59851
   126
	  erls = Rule_Set.Empty,
walther@59851
   127
	  srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@60358
   128
	  rules = [
walther@60358
   129
      \<^rule_thm>\<open>multiply_2nd_order\<close>], 
walther@60358
   130
	  scr = Rule.Empty_Prog});
wneuper@59472
   131
\<close>
jan@42354
   132
wneuper@59472
   133
text \<open>store the rule set for math engine\<close>
wenzelm@60289
   134
rule_set_knowledge
wenzelm@60286
   135
  ansatz_rls = ansatz_rls and
wenzelm@60286
   136
  multiply_ansatz = multiply_ansatz and
wenzelm@60286
   137
  equival_trans = equival_trans
jan@42344
   138
wneuper@59472
   139
subsection \<open>Specification\<close>
jan@42344
   140
neuper@42376
   141
consts
neuper@42376
   142
  decomposedFunction :: "real => una"
neuper@42376
   143
Walther@60502
   144
declare [[check_unique = false]] (*WN120307 REMOVE after editing*)
Walther@60502
   145
wenzelm@60306
   146
problem pbl_simp_rat_partfrac : "partial_fraction/rational/simplification" =
wenzelm@60306
   147
  \<open>Rule_Set.append_rules "empty" Rule_Set.empty [(*for preds in where_ TODO*)]\<close>
Walther@60449
   148
  Method_Ref: "simplification/of_rationals/to_partial_fraction"
wenzelm@60306
   149
  Given: "functionTerm t_t" "solveFor v_v"
wenzelm@60306
   150
  (* TODO: call this sub-problem with appropriate functionTerm: 
wenzelm@60306
   151
    leading coefficient of the denominator is 1: to be checked here! and..
wenzelm@60306
   152
  Where: "((get_numerator t_t) has_degree_in v_v) < 
wenzelm@60306
   153
       ((get_denominator t_t) has_degree_in v_v)" TODO*)
wenzelm@60306
   154
  Find: "decomposedFunction p_p'''"
jan@42354
   155
walther@60154
   156
subsection \<open>MethodC\<close>
wneuper@59585
   157
text \<open>rule set for functions called in the Program\<close>
wneuper@59472
   158
ML \<open>
walther@60358
   159
val srls_partial_fraction = Rule_Def.Repeat {id="srls_partial_fraction", 
walther@60358
   160
  preconds = [],
walther@60358
   161
  rew_ord = ("termlessI",termlessI),
walther@60358
   162
  erls = Rule_Set.append_rules "erls_in_srls_partial_fraction" Rule_Set.empty
walther@60358
   163
    [(*for asm in NTH_CONS ...*)
walther@60358
   164
     \<^rule_eval>\<open>less\<close> (Prog_Expr.eval_equ "#less_"),
walther@60358
   165
     (*2nd NTH_CONS pushes n+-1 into asms*)
walther@60358
   166
     \<^rule_eval>\<open>plus\<close> (**)(eval_binop "#add_")], 
walther@60358
   167
  srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@60358
   168
  rules = [
walther@60358
   169
     \<^rule_thm>\<open>NTH_CONS\<close>,
walther@60358
   170
     \<^rule_eval>\<open>plus\<close> (**)(eval_binop "#add_"),
walther@60358
   171
     \<^rule_thm>\<open>NTH_NIL\<close>,
walther@60358
   172
     \<^rule_eval>\<open>Prog_Expr.lhs\<close> (Prog_Expr.eval_lhs "eval_lhs_"),
walther@60358
   173
     \<^rule_eval>\<open>Prog_Expr.rhs\<close> (Prog_Expr.eval_rhs"eval_rhs_"),
walther@60358
   174
     \<^rule_eval>\<open>Prog_Expr.argument_in\<close> (Prog_Expr.eval_argument_in "Prog_Expr.argument_in"),
walther@60358
   175
     \<^rule_eval>\<open>get_denominator\<close> (eval_get_denominator "#get_denominator"),
walther@60358
   176
     \<^rule_eval>\<open>get_numerator\<close> (eval_get_numerator "#get_numerator"),
walther@60358
   177
     \<^rule_eval>\<open>factors_from_solution\<close> (eval_factors_from_solution "#factors_from_solution")],
walther@60358
   178
  scr = Rule.Empty_Prog};
wneuper@59472
   179
\<close>
neuper@42415
   180
s1210629013@55380
   181
(* current version, error outcommented *)
wneuper@59504
   182
partial_function (tailrec) partial_fraction :: "real \<Rightarrow> real \<Rightarrow> real"
wneuper@59504
   183
  where
wneuper@59504
   184
"partial_fraction f_f zzz =              \<comment> \<open>([1], Frm), 3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z)))\<close>
wneuper@59504
   185
(let f_f = Take f_f;                                                             \<comment> \<open>num_orig = 3\<close>
walther@60260
   186
  num_orig = get_numerator f_f;                  \<comment> \<open>([1], Res), 24 / (-1 + -2 * z + 8 * z  \<up>  2)\<close>
walther@60260
   187
  f_f = (Rewrite_Set ''norm_Rational'') f_f;                \<comment> \<open>denom = -1 + -2 * z + 8 * z  \<up>  2\<close>
walther@60260
   188
  denom = get_denominator f_f;                            \<comment> \<open>equ = -1 + -2 * z + 8 * z  \<up>  2 = 0\<close>
wneuper@59504
   189
  equ = denom = (0::real);
walther@60260
   190
  \<comment> \<open>-----                                  ([2], Pbl), solve (-1 + -2 * z + 8 * z  \<up>  2 = 0, z)\<close>
wneuper@59513
   191
  L_L = SubProblem (''Partial_Fractions'',                 \<comment> \<open>([2], Res), [z = 1 / 2, z = -1 / 4\<close>
wneuper@59504
   192
    [''abcFormula'', ''degree_2'', ''polynomial'', ''univariate'', ''equation''],
wneuper@59504
   193
    [''no_met'']) [BOOL equ, REAL zzz];                      \<comment> \<open>facs: (z - 1 / 2) * (z - -1 / 4)\<close>
wneuper@59504
   194
  facs = factors_from_solution L_L;             \<comment> \<open>([3], Frm), 33 / ((z - 1 / 2) * (z - -1 / 4))\<close>
wneuper@59504
   195
  eql = Take (num_orig / facs);              \<comment> \<open>([3], Res), ?A / (z - 1 / 2) + ?B / (z - -1 / 4)\<close>
walther@59635
   196
  eqr = (Try (Rewrite_Set ''ansatz_rls'')) eql;
wneuper@59504
   197
          \<comment> \<open>([4], Frm), 3 / ((z - 1 / 2) * (z - -1 / 4)) = ?A / (z - 1 / 2) + ?B / (z - -1 / 4)\<close>
wneuper@59536
   198
  eq = Take (eql = eqr);                 \<comment> \<open>([4], Res), 3 = ?A * (z - -1 / 4) + ?B * (z - 1 / 2)\<close>
walther@59635
   199
  eq = (Try (Rewrite_Set ''equival_trans'')) eq; 
wneuper@59536
   200
                                                \<comment> \<open>eq = 3 = AA * (z - -1 / 4) + BB * (z - 1 / 2)\<close>
wneuper@59504
   201
  z1 = rhs (NTH 1 L_L);                                                           \<comment> \<open>z2 = -1 / 4\<close>
wneuper@59536
   202
  z2 = rhs (NTH 2 L_L);                  \<comment> \<open>([5], Frm), 3 = AA * (z - -1 / 4) + BB * (z - 1 / 2)\<close>
wneuper@59536
   203
  eq_a = Take eq;                \<comment> \<open>([5], Res), 3 = AA * (1 / 2 - -1 / 4) + BB * (1 / 2 - 1 / 2)\<close>
wneuper@59536
   204
  eq_a = Substitute [zzz = z1] eq;                                 \<comment> \<open>([6], Res), 3 = 3 * AA / 4\<close>
walther@59635
   205
  eq_a = (Rewrite_Set ''norm_Rational'') eq_a;
wneuper@59536
   206
\<comment> \<open>-----                                                  ([7], Pbl), solve (3 = 3 * AA / 4, AA)\<close>
wneuper@59536
   207
                                                                         \<comment> \<open>([7], Res), [AA = 4]\<close>
wneuper@59592
   208
  sol_a = SubProblem (''Isac_Knowledge'', [''univariate'',''equation''], [''no_met''])
wneuper@59536
   209
      [BOOL eq_a, REAL (AA::real)] ;                                                    \<comment> \<open>a = 4\<close>
wneuper@59536
   210
  a = rhs (NTH 1 sol_a);                 \<comment> \<open>([8], Frm), 3 = AA * (z - -1 / 4) + BB * (z - 1 / 2)\<close>
wneuper@59536
   211
  eq_b = Take eq;              \<comment> \<open>([8], Res), 3 = AA * (-1 / 4 - -1 / 4) + BB * (-1 / 4 - 1 / 2)\<close>
wneuper@59536
   212
  eq_b = Substitute [zzz = z2] eq_b;                              \<comment> \<open>([9], Res), 3 = -3 * BB / 4\<close>
walther@59635
   213
  eq_b = (Rewrite_Set ''norm_Rational'') eq_b;       \<comment> \<open>([10], Pbl), solve (3 = -3 * BB / 4, BB)\<close>
walther@59635
   214
  sol_b = SubProblem (''Isac_Knowledge'',                              \<comment> \<open>([10], Res), [BB = -4]\<close>
wneuper@59504
   215
      [''univariate'',''equation''], [''no_met''])
wneuper@59536
   216
    [BOOL eq_b, REAL (BB::real)];                                                      \<comment> \<open>b = -4\<close>
wneuper@59536
   217
  b = rhs (NTH 1 sol_b);                           \<comment> \<open>eqr = AA / (z - 1 / 2) + BB / (z - -1 / 4)\<close>
wneuper@59504
   218
  pbz = Take eqr;                            \<comment> \<open>([11], Res), 4 / (z - 1 / 2) + -4 / (z - -1 / 4)\<close>
wneuper@59512
   219
  pbz = Substitute [AA = a, BB = b] pbz        \<comment> \<open>([], Res), 4 / (z - 1 / 2) + -4 / (z - -1 / 4)\<close>
wneuper@59504
   220
in pbz)                                                                                "
wenzelm@60303
   221
wenzelm@60303
   222
method met_partial_fraction : "simplification/of_rationals/to_partial_fraction" =
wenzelm@60303
   223
  \<open>{rew_ord'="tless_true", rls'= Rule_Set.empty, calc = [], srls = srls_partial_fraction, prls = Rule_Set.empty,
wenzelm@60303
   224
   crls = Rule_Set.empty, errpats = [], nrls = Rule_Set.empty}\<close>
wenzelm@60303
   225
    (*f_f = 3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z)), zzz: z*)
wenzelm@60303
   226
    (*([], Frm), Problem (Partial_Fractions, [partial_fraction, rational, simplification])*)
wenzelm@60303
   227
  Program: partial_fraction.simps
wenzelm@60303
   228
  Given: "functionTerm t_t" "solveFor v_v"
wenzelm@60303
   229
  (* Where: "((get_numerator t_t) has_degree_in v_v) <
wenzelm@60303
   230
    ((get_denominator t_t) has_degree_in v_v)" TODO *)
wenzelm@60303
   231
  Find: "decomposedFunction p_p'''"
wenzelm@60303
   232
wneuper@59472
   233
ML \<open>
neuper@42376
   234
(*
neuper@42376
   235
  val fmz =                                             
neuper@42376
   236
    ["functionTerm (3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z))))", 
neuper@42376
   237
      "solveFor z", "functionTerm p_p"];
neuper@42376
   238
  val (dI',pI',mI') =
neuper@42376
   239
    ("Partial_Fractions", 
neuper@42376
   240
      ["partial_fraction", "rational", "simplification"],
walther@59997
   241
      ["simplification", "of_rationals", "to_partial_fraction"]);
neuper@42376
   242
  val (p,_,f,nxt,_,pt) = CalcTreeTEST [(fmz, (dI',pI',mI'))];
neuper@42376
   243
*)
walther@60278
   244
\<close>ML \<open>
walther@60278
   245
\<close> ML \<open>
wneuper@59472
   246
\<close>
neuper@42289
   247
end