src/Tools/isac/Knowledge/Partial_Fractions.thy
author Walther Neuper <wneuper@ist.tugraz.at>
Thu, 07 Mar 2019 17:22:20 +0100
changeset 59512 e504168e7b01
parent 59505 a1f223658994
child 59513 deb1efba3119
permissions -rw-r--r--
[-Test_Isac] funpack: Const ("Partial_Fractions.AA",..) makes trick superfluous
     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 "Groups.times_class.times" (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 "Groups.minus_class.minus" (lhs, rhs) end;
    40 
    41 fun mk_prod prod [] =
    42       if prod = Rule.e_term then error "mk_prod called with []" else prod
    43   | mk_prod prod (t :: []) =
    44       if prod = Rule.e_term then t else HOLogic.mk_binop "Groups.times_class.times" (prod, t)
    45   | mk_prod prod (t1 :: t2 :: ts) =
    46         if prod = Rule.e_term 
    47         then 
    48            let val p = HOLogic.mk_binop "Groups.times_class.times" (t1, t2)
    49            in mk_prod p ts end 
    50         else 
    51            let val p = HOLogic.mk_binop "Groups.times_class.times" (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 Rule.e_term (map fac_from_sol ts) end;
    57 
    58 (*("factors_from_solution", ("Partial_Fractions.factors_from_solution", 
    59      eval_factors_from_solution ""))*)
    60 fun eval_factors_from_solution (thmid:string) _
    61      (t as Const ("Partial_Fractions.factors_from_solution", _) $ sol) thy =
    62        ((let val prod = factors_from_solution sol
    63          in SOME (TermC.mk_thmid thmid (Rule.term_to_string''' thy prod) "",
    64               HOLogic.Trueprop $ (TermC.mk_equality (t, prod)))
    65          end)
    66        handle _ => NONE)
    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 ar in partial_fractions.sml "--- progr.vers.2: "
   105 \<close>
   106 
   107 ML \<open>
   108 val ansatz_rls = prep_rls'(
   109   Rule.Rls {id = "ansatz_rls", preconds = [], rew_ord = ("dummy_ord",Rule.dummy_ord), 
   110 	  erls = Rule.Erls, srls = Rule.Erls, calc = [], errpatts = [],
   111 	  rules = 
   112 	   [Rule.Thm ("ansatz_2nd_order",TermC.num_str @{thm ansatz_2nd_order}),
   113 	    Rule.Thm ("ansatz_3rd_order",TermC.num_str @{thm ansatz_3rd_order})
   114 	   ], 
   115 	 scr = Rule.EmptyScr});
   116 
   117 val equival_trans = prep_rls'(
   118   Rule.Rls {id = "equival_trans", preconds = [], rew_ord = ("dummy_ord",Rule.dummy_ord), 
   119 	  erls = Rule.Erls, srls = Rule.Erls, calc = [], errpatts = [],
   120 	  rules = 
   121 	   [Rule.Thm ("equival_trans_2nd_order",TermC.num_str @{thm equival_trans_2nd_order}),
   122 	    Rule.Thm ("equival_trans_3rd_order",TermC.num_str @{thm equival_trans_3rd_order})
   123 	   ], 
   124 	 scr = Rule.EmptyScr});
   125 
   126 val multiply_ansatz = prep_rls'(
   127   Rule.Rls {id = "multiply_ansatz", preconds = [], rew_ord = ("dummy_ord",Rule.dummy_ord), 
   128 	  erls = Rule.Erls,
   129 	  srls = Rule.Erls, calc = [], errpatts = [],
   130 	  rules = 
   131 	   [Rule.Thm ("multiply_2nd_order",TermC.num_str @{thm multiply_2nd_order})
   132 	   ], 
   133 	 scr = Rule.EmptyScr});
   134 \<close>
   135 
   136 text \<open>store the rule set for math engine\<close>
   137 setup \<open>KEStore_Elems.add_rlss 
   138   [("ansatz_rls", (Context.theory_name @{theory}, ansatz_rls)), 
   139   ("multiply_ansatz", (Context.theory_name @{theory}, multiply_ansatz)), 
   140   ("equival_trans", (Context.theory_name @{theory}, equival_trans))]\<close>
   141 
   142 subsection \<open>Specification\<close>
   143 
   144 consts
   145   decomposedFunction :: "real => una"
   146 
   147 ML \<open>
   148 Celem.check_guhs_unique := false; (*WN120307 REMOVE after editing*)
   149 \<close>
   150 setup \<open>KEStore_Elems.add_pbts
   151   [(Specify.prep_pbt @{theory} "pbl_simp_rat_partfrac" [] Celem.e_pblID
   152       (["partial_fraction", "rational", "simplification"],
   153         [("#Given" ,["functionTerm t_t", "solveFor v_v"]),
   154           (* TODO: call this sub-problem with appropriate functionTerm: 
   155             leading coefficient of the denominator is 1: to be checked here! and..
   156             ("#Where" ,["((get_numerator t_t) has_degree_in v_v) < 
   157                ((get_denominator t_t) has_degree_in v_v)"]), TODO*)
   158           ("#Find"  ,["decomposedFunction p_p'''"])],
   159         Rule.append_rls "e_rls" Rule.e_rls [(*for preds in where_ TODO*)], 
   160         NONE, 
   161         [["simplification","of_rationals","to_partial_fraction"]]))]\<close>
   162 
   163 subsection \<open>Method\<close>
   164 consts
   165   PartFracScript  :: "[real,real,  real] => real" 
   166     ("((Script PartFracScript (_ _ =))// (_))" 9)
   167 
   168 text \<open>rule set for functions called in the Script\<close>
   169 ML \<open>
   170   val srls_partial_fraction = Rule.Rls {id="srls_partial_fraction", 
   171     preconds = [],
   172     rew_ord = ("termlessI",termlessI),
   173     erls = Rule.append_rls "erls_in_srls_partial_fraction" Rule.e_rls
   174       [(*for asm in NTH_CONS ...*)
   175        Rule.Calc ("Orderings.ord_class.less",eval_equ "#less_"),
   176        (*2nd NTH_CONS pushes n+-1 into asms*)
   177        Rule.Calc("Groups.plus_class.plus", eval_binop "#add_")], 
   178     srls = Rule.Erls, calc = [], errpatts = [],
   179     rules = [
   180        Rule.Thm ("NTH_CONS",TermC.num_str @{thm NTH_CONS}),
   181        Rule.Calc("Groups.plus_class.plus", eval_binop "#add_"),
   182        Rule.Thm ("NTH_NIL",TermC.num_str @{thm NTH_NIL}),
   183        Rule.Calc("Tools.lhs", Tools.eval_lhs "eval_lhs_"),
   184        Rule.Calc("Tools.rhs", Tools.eval_rhs"eval_rhs_"),
   185        Rule.Calc("Atools.argument'_in", eval_argument_in "Atools.argument'_in"),
   186        Rule.Calc("Rational.get_denominator", eval_get_denominator "#get_denominator"),
   187        Rule.Calc("Rational.get_numerator", eval_get_numerator "#get_numerator"),
   188        Rule.Calc("Partial_Fractions.factors_from_solution",
   189          eval_factors_from_solution "#factors_from_solution")
   190        ],
   191     scr = Rule.EmptyScr};
   192 \<close>
   193 
   194 (* current version, error outcommented *)
   195 (*ok
   196 partial_function (tailrec) partial_fraction :: "real \<Rightarrow> real \<Rightarrow> real"
   197   where
   198 "partial_fraction f_f zzz =              \<comment> \<open>([1], Frm), 3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z)))\<close>
   199 (let f_f = Take f_f;                                                             \<comment> \<open>num_orig = 3\<close>
   200   num_orig = get_numerator f_f;                  \<comment> \<open>([1], Res), 24 / (-1 + -2 * z + 8 * z ^^^ 2)\<close>
   201   f_f = (Rewrite_Set ''norm_Rational'' False) f_f;          \<comment> \<open>denom = -1 + -2 * z + 8 * z ^^^ 2\<close>
   202   denom = get_denominator f_f;                            \<comment> \<open>equ = -1 + -2 * z + 8 * z ^^^ 2 = 0\<close>
   203   equ = denom = (0::real);
   204   \<comment> \<open>-----                                  ([2], Pbl), solve (-1 + -2 * z + 8 * z ^^^ 2 = 0, z)\<close>
   205   L_L = SubProblem (''PolyEq'',                            \<comment> \<open>([2], Res), [z = 1 / 2, z = -1 / 4\<close>
   206     [''abcFormula'', ''degree_2'', ''polynomial'', ''univariate'', ''equation''],
   207     [''no_met'']) [BOOL equ, REAL zzz];                      \<comment> \<open>facs: (z - 1 / 2) * (z - -1 / 4)\<close>
   208   facs = factors_from_solution L_L;             \<comment> \<open>([3], Frm), 33 / ((z - 1 / 2) * (z - -1 / 4))\<close>
   209   eql = Take (num_orig / facs);              \<comment> \<open>([3], Res), ?A / (z - 1 / 2) + ?B / (z - -1 / 4)\<close>
   210   eqr = (Try (Rewrite_Set ''ansatz_rls'' False)) eql;
   211           \<comment> \<open>([4], Frm), 3 / ((z - 1 / 2) * (z - -1 / 4)) = ?A / (z - 1 / 2) + ?B / (z - -1 / 4)\<close>
   212   eq = Take (eql = eqr);                  \<comment> \<open>([4], Res), 3 = ?A * (z - -1 / 4) + ?B * (z - 1 / 2)\<close>
   213   eq = (Try (Rewrite_Set ''equival_trans'' False)) eq; 
   214                                                  \<comment> \<open>eq = 3 = AA * (z - -1 / 4) + BB * (z - 1 / 2)\<close>
   215   z1 = rhs (NTH 1 L_L);                                                           \<comment> \<open>z2 = -1 / 4\<close>
   216   z2 = rhs (NTH 2 L_L);                   \<comment> \<open>([5], Frm), 3 = AA * (z - -1 / 4) + BB * (z - 1 / 2)\<close>
   217   eq_a = Take eq;                 \<comment> \<open>([5], Res), 3 = AA * (1 / 2 - -1 / 4) + BB * (1 / 2 - 1 / 2)\<close>
   218   eq_a = Substitute [zzz = z1] eq;                                \<comment> \<open>([6], Res), 3 = 3 * AA / 4\<close>
   219   eq_a = (Rewrite_Set ''norm_Rational'' False) eq_a;
   220 \<comment> \<open>-----                                                    ([7], Pbl), solve (3 = 3 * AA / 4, AA)\<close>
   221                                                                           \<comment> \<open>([7], Res), [AA = 4]\<close>
   222   sol_a = SubProblem (''Isac'', [''univariate'',''equation''], [''no_met''])
   223       [BOOL eq_a, REAL (AA::real)] ;                                                     \<comment> \<open>a = 4\<close>
   224   a = rhs (NTH 1 sol_a);                   \<comment> \<open>([8], Frm), 3 = AA * (z - -1 / 4) + BB * (z - 1 / 2)\<close>
   225   eq_b = Take eq;                \<comment> \<open>([8], Res), 3 = AA * (-1 / 4 - -1 / 4) + BB * (-1 / 4 - 1 / 2)\<close>
   226   eq_b = Substitute [zzz = z2] eq_b;                               \<comment> \<open>([9], Res), 3 = -3 * BB / 4\<close>
   227   eq_b = (Rewrite_Set ''norm_Rational'' False) eq_b;   \<comment> \<open>([10], Pbl), solve (3 = -3 * BB / 4, BB)\<close>
   228   sol_b = SubProblem (''Isac'',                                         \<comment> \<open>([10], Res), [BB = -4]\<close>
   229       [''univariate'',''equation''], [''no_met''])
   230     [BOOL eq_b, REAL (BB::real)];                                                       \<comment> \<open>b = -4\<close>
   231   b = rhs (NTH 1 sol_b);                             \<comment> \<open>eqr = AA / (z - 1 / 2) + BB / (z - -1 / 4)\<close>
   232   pbz = Take eqr;                            \<comment> \<open>([11], Res), 4 / (z - 1 / 2) + -4 / (z - -1 / 4)\<close>
   233   pbz = Substitute [AA = a, BB = b] pbz        \<comment> \<open>([], Res), 4 / (z - 1 / 2) + -4 / (z - -1 / 4)\<close>
   234 in pbz)                                                                                "
   235 *)
   236 setup \<open>KEStore_Elems.add_mets
   237     [Specify.prep_met @{theory} "met_partial_fraction" [] Celem.e_metID
   238       (["simplification","of_rationals","to_partial_fraction"], 
   239         [("#Given" ,["functionTerm t_t", "solveFor v_v"]),
   240           (*("#Where" ,["((get_numerator t_t) has_degree_in v_v) < 
   241             ((get_denominator t_t) has_degree_in v_v)"]), TODO*)
   242           ("#Find"  ,["decomposedFunction p_p'''"])],
   243         (*f_f = 3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z)), zzz: z*)
   244         {rew_ord'="tless_true", rls'= Rule.e_rls, calc = [], srls = srls_partial_fraction, prls = Rule.e_rls,
   245           crls = Rule.e_rls, errpats = [], nrls = Rule.e_rls},
   246         (*([], Frm), Problem (Partial_Fractions, [partial_fraction, rational, simplification])*)
   247         "Script PartFracScript (f_f::real) (zzz::real) =   " ^
   248           (*([1], Frm), 3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z)))*)
   249           "(let f_f = Take f_f;                              " ^
   250           (*           num_orig = 3*)
   251           "  (num_orig::real) = get_numerator f_f;           " ^
   252           (*([1], Res), 24 / (-1 + -2 * z + 8 * z ^^^ 2)*)
   253           "  f_f = (Rewrite_Set ''norm_Rational'' False) f_f;    " ^
   254           (*           denom = -1 + -2 * z + 8 * z ^^^ 2*)
   255           "  (denom::real) = get_denominator f_f;            " ^
   256           (*           equ = -1 + -2 * z + 8 * z ^^^ 2 = 0*)
   257           "  (equ::bool) = (denom = (0::real));              " ^
   258 
   259           (*([2], Pbl), solve (-1 + -2 * z + 8 * z ^^^ 2 = 0, z)*)
   260           "  (L_L::bool list) = (SubProblem (''PolyEq'',        " ^
   261           "    [''abcFormula'', ''degree_2'', ''polynomial'', ''univariate'', ''equation''], " ^
   262           (*([2], Res), [z = 1 / 2, z = -1 / 4]*)
   263           "    [''no_met'']) [BOOL equ, REAL zzz]);              " ^
   264           (*           facs: (z - 1 / 2) * (z - -1 / 4)*)
   265           "  (facs::real) = factors_from_solution L_L;       " ^
   266           (*([3], Frm), 33 / ((z - 1 / 2) * (z - -1 / 4)) *) 
   267           "  (eql::real) = Take (num_orig / facs);           " ^
   268           (*([3], Res), ?A / (z - 1 / 2) + ?B / (z - -1 / 4)*)
   269           "  (eqr::real) = (Try (Rewrite_Set ''ansatz_rls'' False)) eql;  " ^
   270           (*([4], Frm), 3 / ((z - 1 / 2) * (z - -1 / 4)) = ?A / (z - 1 / 2) + ?B / (z - -1 / 4)*)
   271           "  (eq::bool) = Take (eql = eqr);                  " ^
   272           (*([4], Res), 3 = ?A * (z - -1 / 4) + ?B * (z - 1 / 2)*)
   273           "  eq = (Try (Rewrite_Set ''equival_trans'' False)) eq;" ^
   274           (*           eq = 3 = AA * (z - -1 / 4) + BB * (z - 1 / 2)*)
   275           (*           z1 = 1 / 2*)
   276           "  (z1::real) = (rhs (NTH 1 L_L));                 " ^
   277           (*           z2 = -1 / 4*)
   278           "  (z2::real) = (rhs (NTH 2 L_L));                 " ^
   279           (*([5], Frm), 3 = AA * (z - -1 / 4) + BB * (z - 1 / 2)*)
   280           "  (eq_a::bool) = Take eq;                         " ^
   281           (*([5], Res), 3 = AA * (1 / 2 - -1 / 4) + BB * (1 / 2 - 1 / 2)*)
   282           "  eq_a = (Substitute [zzz = z1]) eq;              " ^
   283           (*([6], Res), 3 = 3 * AA / 4*)
   284           "  eq_a = (Rewrite_Set ''norm_Rational'' False) eq_a;  " ^
   285 
   286           (*([7], Pbl), solve (3 = 3 * AA / 4, AA)*)
   287           "  (sol_a::bool list) =                            " ^
   288           "    (SubProblem (''Isac'', [''univariate'',''equation''], [''no_met''])   " ^
   289           (*([7], Res), [AA = 4]*)
   290           "    [BOOL eq_a, REAL (AA::real)]);                 " ^
   291           (*           a = 4*)
   292           "  (a::real) = (rhs (NTH 1 sol_a));                " ^
   293           (*([8], Frm), 3 = AA * (z - -1 / 4) + BB * (z - 1 / 2)*)
   294           "  (eq_b::bool) = Take eq;                         " ^
   295           (*([8], Res), 3 = AA * (-1 / 4 - -1 / 4) + BB * (-1 / 4 - 1 / 2)*)
   296           "  eq_b = (Substitute [zzz = z2]) eq_b;            " ^
   297           (*([9], Res), 3 = -3 * BB / 4*)
   298           "  eq_b = (Rewrite_Set ''norm_Rational'' False) eq_b;  " ^
   299           (*([10], Pbl), solve (3 = -3 * BB / 4, BB)*)
   300           "  (sol_b::bool list) =                            " ^
   301           "    (SubProblem (''Isac'', [''univariate'',''equation''], [''no_met''])   " ^
   302           (*([10], Res), [BB = -4]*)
   303           "    [BOOL eq_b, REAL (BB::real)]);                 " ^
   304           (*           b = -4*)
   305           "  (b::real) = (rhs (NTH 1 sol_b));                " ^
   306           (*           eqr = AA / (z - 1 / 2) + BB / (z - -1 / 4)*)
   307           (*([11], Frm), AA / (z - 1 / 2) + BB / (z - -1 / 4)*)
   308           "  (pbz::real) = Take eqr;                         " ^
   309           (*([11], Res), 4 / (z - 1 / 2) + -4 / (z - -1 / 4)*)
   310           "  pbz = ((Substitute [AA = a, BB = b]) pbz)         " ^
   311           (*([], Res), 4 / (z - 1 / 2) + -4 / (z - -1 / 4)*)
   312           "in pbz)"
   313 )]
   314 \<close>
   315 ML \<open>
   316 (*
   317   val fmz =                                             
   318     ["functionTerm (3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z))))", 
   319       "solveFor z", "functionTerm p_p"];
   320   val (dI',pI',mI') =
   321     ("Partial_Fractions", 
   322       ["partial_fraction", "rational", "simplification"],
   323       ["simplification","of_rationals","to_partial_fraction"]);
   324   val (p,_,f,nxt,_,pt) = CalcTreeTEST [(fmz, (dI',pI',mI'))];
   325 *)
   326 \<close>
   327 
   328 end