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