neuper@42376: (* Partial_Fractions jan@42344: author: Jan Rocnik, isac team jan@42344: Copyright (c) isac team 2011 jan@42344: Use is subject to license terms. jan@42344: *) wneuper@59344: (* Partial Fraction Decomposition *) jan@42344: neuper@42289: neuper@42376: theory Partial_Fractions imports RootRatEq begin jan@42353: wneuper@59472: ML\ neuper@42376: (* jan@42353: signature PARTFRAC = jan@42353: sig jan@42353: val ansatz_rls : rls jan@42353: val ansatz_rls_ : theory -> term -> (term * term list) option jan@42353: end neuper@42376: *) wneuper@59472: \ jan@42353: wneuper@59472: subsection \eval_ functions\ jan@42344: consts neuper@42359: factors_from_solution :: "bool list => real" wneuper@59512: AA :: real wneuper@59512: BB :: real jan@42344: wneuper@59472: text \these might be used for variants of fac_from_sol\ wneuper@59472: ML \ jan@42344: fun mk_minus_1 T = Free("-1", T); (*TODO DELETE WITH numbers_to_string*) jan@42344: fun flip_sign t = (*TODO improve for use in factors_from_solution: -(-1) etc*) jan@42344: let val minus_1 = t |> type_of |> mk_minus_1 wenzelm@60309: in HOLogic.mk_binop \<^const_name>\times\ (minus_1, t) end; wneuper@59472: \ neuper@42376: wneuper@59472: text \from solutions (e.g. [z = 1, z = -2]) make linear factors (e.g. (z - 1)*(z - -2))\ wneuper@59472: ML \ jan@42344: fun fac_from_sol s = jan@42344: let val (lhs, rhs) = HOLogic.dest_eq s wenzelm@60309: in HOLogic.mk_binop \<^const_name>\minus\ (lhs, rhs) end; jan@42344: jan@42344: fun mk_prod prod [] = walther@59962: if prod = TermC.empty then raise ERROR "mk_prod called with []" else prod jan@42344: | mk_prod prod (t :: []) = wenzelm@60309: if prod = TermC.empty then t else HOLogic.mk_binop \<^const_name>\times\ (prod, t) jan@42344: | mk_prod prod (t1 :: t2 :: ts) = walther@59861: if prod = TermC.empty jan@42344: then wenzelm@60309: let val p = HOLogic.mk_binop \<^const_name>\times\ (t1, t2) jan@42344: in mk_prod p ts end jan@42344: else wenzelm@60309: let val p = HOLogic.mk_binop \<^const_name>\times\ (prod, t1) jan@42344: in mk_prod p (t2 :: ts) end jan@42344: jan@42344: fun factors_from_solution sol = jan@42344: let val ts = HOLogic.dest_list sol walther@59861: in mk_prod TermC.empty (map fac_from_sol ts) end; jan@42344: neuper@42376: (*("factors_from_solution", ("Partial_Fractions.factors_from_solution", walther@60264: eval_factors_from_solution "")) walther@60264: this code is limited (max 3 solutions) AND has too little checks *) jan@42352: fun eval_factors_from_solution (thmid:string) _ Walther@60504: (t as Const (\<^const_name>\Partial_Fractions.factors_from_solution\, _) $ sol) ctxt = walther@60264: let val prod = factors_from_solution sol Walther@60504: in SOME (TermC.mk_thmid thmid (UnparseC.term_in_ctxt ctxt prod) "", walther@60264: HOLogic.Trueprop $ (TermC.mk_equality (t, prod))) walther@60264: end walther@60264: | eval_factors_from_solution _ _ _ _ = NONE; wneuper@59472: \ jan@42344: wneuper@59472: subsection \'ansatz' for partial fractions\ jan@42353: axiomatization where wneuper@59512: ansatz_2nd_order: "n / (a*b) = AA/a + BB/b" and wneuper@59512: ansatz_3rd_order: "n / (a*b*c) = AA/a + BB/b + C/c" and wneuper@59512: ansatz_4th_order: "n / (a*b*c*d) = AA/a + BB/b + C/c + D/d" and neuper@42386: (*version 1*) wneuper@59512: equival_trans_2nd_order: "(n/(a*b) = AA/a + BB/b) = (n = AA*b + BB*a)" and wneuper@59512: 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: equival_trans_4th_order: "(n/(a*b*c*d) = AA/a + BB/b + C/c + D/d) = wneuper@59512: (n = AA*b*c*d + BB*a*c*d + C*a*b*d + D*a*b*c)" and neuper@42386: (*version 2: not yet used, see partial_fractions.sml*) wneuper@59512: multiply_2nd_order: "(n/x = AA/a + BB/b) = (a*b*n/x = AA*b + BB*a)" and wneuper@59512: 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: multiply_4th_order: wneuper@59512: "(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: wneuper@59472: text \Probably the optimal formalization woudl be ... neuper@42387: wneuper@59512: multiply_2nd_order: "x = a*b ==> (n/x = AA/a + BB/b) = (a*b*n/x = AA*b + BB*a)" and neuper@42386: multiply_3rd_order: "x = a*b*c ==> wneuper@59512: (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: multiply_4th_order: "x = a*b*c*d ==> wneuper@59512: (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: neuper@42387: ... because it would allow to start the ansatz as follows neuper@42387: (1) 3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z))) = 3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z))) neuper@42387: (2) 3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z))) = AA / (z - 1 / 2) + BB / (z - -1 / 4) neuper@42387: (3) (z - 1 / 2) * (z - -1 / 4) * 3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z))) = neuper@42387: (z - 1 / 2) * (z - -1 / 4) * AA / (z - 1 / 2) + BB / (z - -1 / 4) wneuper@59512: (4) 3 = AA * (z - -1 / 4) + BB * (z - 1 / 2) neuper@42387: neuper@42387: ... (1==>2) ansatz neuper@42387: (2==>3) multiply_* neuper@42387: (3==>4) norm_Rational wneuper@59550: TODOs for this version are in partial_fractions.sml "--- progr.vers.2: " wneuper@59472: \ neuper@42387: wneuper@59472: ML \ s1210629013@55444: val ansatz_rls = prep_rls'( Walther@60509: Rule_Def.Repeat {id = "ansatz_rls", preconds = [], rew_ord = ("dummy_ord",Rewrite_Ord.function_empty), walther@59851: erls = Rule_Set.Empty, srls = Rule_Set.Empty, calc = [], errpatts = [], walther@60358: rules = [ walther@60358: \<^rule_thm>\ansatz_2nd_order\, walther@60358: \<^rule_thm>\ansatz_3rd_order\], walther@60358: scr = Rule.Empty_Prog}); jan@42353: s1210629013@55444: val equival_trans = prep_rls'( Walther@60509: Rule_Def.Repeat {id = "equival_trans", preconds = [], rew_ord = ("dummy_ord",Rewrite_Ord.function_empty), walther@59851: erls = Rule_Set.Empty, srls = Rule_Set.Empty, calc = [], errpatts = [], walther@60358: rules = [ walther@60358: \<^rule_thm>\equival_trans_2nd_order\, walther@60358: \<^rule_thm>\equival_trans_3rd_order\], walther@60358: scr = Rule.Empty_Prog}); neuper@42386: s1210629013@55444: val multiply_ansatz = prep_rls'( Walther@60509: Rule_Def.Repeat {id = "multiply_ansatz", preconds = [], rew_ord = ("dummy_ord",Rewrite_Ord.function_empty), walther@59851: erls = Rule_Set.Empty, walther@59851: srls = Rule_Set.Empty, calc = [], errpatts = [], walther@60358: rules = [ walther@60358: \<^rule_thm>\multiply_2nd_order\], walther@60358: scr = Rule.Empty_Prog}); wneuper@59472: \ jan@42354: wneuper@59472: text \store the rule set for math engine\ wenzelm@60289: rule_set_knowledge wenzelm@60286: ansatz_rls = ansatz_rls and wenzelm@60286: multiply_ansatz = multiply_ansatz and wenzelm@60286: equival_trans = equival_trans jan@42344: wneuper@59472: subsection \Specification\ jan@42344: neuper@42376: consts neuper@42376: decomposedFunction :: "real => una" neuper@42376: Walther@60502: declare [[check_unique = false]] (*WN120307 REMOVE after editing*) Walther@60502: wenzelm@60306: problem pbl_simp_rat_partfrac : "partial_fraction/rational/simplification" = wenzelm@60306: \Rule_Set.append_rules "empty" Rule_Set.empty [(*for preds in where_ TODO*)]\ Walther@60449: Method_Ref: "simplification/of_rationals/to_partial_fraction" wenzelm@60306: Given: "functionTerm t_t" "solveFor v_v" wenzelm@60306: (* TODO: call this sub-problem with appropriate functionTerm: wenzelm@60306: leading coefficient of the denominator is 1: to be checked here! and.. wenzelm@60306: Where: "((get_numerator t_t) has_degree_in v_v) < wenzelm@60306: ((get_denominator t_t) has_degree_in v_v)" TODO*) wenzelm@60306: Find: "decomposedFunction p_p'''" jan@42354: walther@60154: subsection \MethodC\ wneuper@59585: text \rule set for functions called in the Program\ wneuper@59472: ML \ walther@60358: val srls_partial_fraction = Rule_Def.Repeat {id="srls_partial_fraction", walther@60358: preconds = [], walther@60358: rew_ord = ("termlessI",termlessI), walther@60358: erls = Rule_Set.append_rules "erls_in_srls_partial_fraction" Rule_Set.empty walther@60358: [(*for asm in NTH_CONS ...*) walther@60358: \<^rule_eval>\less\ (Prog_Expr.eval_equ "#less_"), walther@60358: (*2nd NTH_CONS pushes n+-1 into asms*) walther@60358: \<^rule_eval>\plus\ (**)(eval_binop "#add_")], walther@60358: srls = Rule_Set.Empty, calc = [], errpatts = [], walther@60358: rules = [ walther@60358: \<^rule_thm>\NTH_CONS\, walther@60358: \<^rule_eval>\plus\ (**)(eval_binop "#add_"), walther@60358: \<^rule_thm>\NTH_NIL\, walther@60358: \<^rule_eval>\Prog_Expr.lhs\ (Prog_Expr.eval_lhs "eval_lhs_"), walther@60358: \<^rule_eval>\Prog_Expr.rhs\ (Prog_Expr.eval_rhs"eval_rhs_"), walther@60358: \<^rule_eval>\Prog_Expr.argument_in\ (Prog_Expr.eval_argument_in "Prog_Expr.argument_in"), walther@60358: \<^rule_eval>\get_denominator\ (eval_get_denominator "#get_denominator"), walther@60358: \<^rule_eval>\get_numerator\ (eval_get_numerator "#get_numerator"), walther@60358: \<^rule_eval>\factors_from_solution\ (eval_factors_from_solution "#factors_from_solution")], walther@60358: scr = Rule.Empty_Prog}; wneuper@59472: \ neuper@42415: s1210629013@55380: (* current version, error outcommented *) wneuper@59504: partial_function (tailrec) partial_fraction :: "real \ real \ real" wneuper@59504: where wneuper@59504: "partial_fraction f_f zzz = \ \([1], Frm), 3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z)))\ wneuper@59504: (let f_f = Take f_f; \ \num_orig = 3\ walther@60260: num_orig = get_numerator f_f; \ \([1], Res), 24 / (-1 + -2 * z + 8 * z \ 2)\ walther@60260: f_f = (Rewrite_Set ''norm_Rational'') f_f; \ \denom = -1 + -2 * z + 8 * z \ 2\ walther@60260: denom = get_denominator f_f; \ \equ = -1 + -2 * z + 8 * z \ 2 = 0\ wneuper@59504: equ = denom = (0::real); walther@60260: \ \----- ([2], Pbl), solve (-1 + -2 * z + 8 * z \ 2 = 0, z)\ wneuper@59513: L_L = SubProblem (''Partial_Fractions'', \ \([2], Res), [z = 1 / 2, z = -1 / 4\ wneuper@59504: [''abcFormula'', ''degree_2'', ''polynomial'', ''univariate'', ''equation''], wneuper@59504: [''no_met'']) [BOOL equ, REAL zzz]; \ \facs: (z - 1 / 2) * (z - -1 / 4)\ wneuper@59504: facs = factors_from_solution L_L; \ \([3], Frm), 33 / ((z - 1 / 2) * (z - -1 / 4))\ wneuper@59504: eql = Take (num_orig / facs); \ \([3], Res), ?A / (z - 1 / 2) + ?B / (z - -1 / 4)\ walther@59635: eqr = (Try (Rewrite_Set ''ansatz_rls'')) eql; wneuper@59504: \ \([4], Frm), 3 / ((z - 1 / 2) * (z - -1 / 4)) = ?A / (z - 1 / 2) + ?B / (z - -1 / 4)\ wneuper@59536: eq = Take (eql = eqr); \ \([4], Res), 3 = ?A * (z - -1 / 4) + ?B * (z - 1 / 2)\ walther@59635: eq = (Try (Rewrite_Set ''equival_trans'')) eq; wneuper@59536: \ \eq = 3 = AA * (z - -1 / 4) + BB * (z - 1 / 2)\ wneuper@59504: z1 = rhs (NTH 1 L_L); \ \z2 = -1 / 4\ wneuper@59536: z2 = rhs (NTH 2 L_L); \ \([5], Frm), 3 = AA * (z - -1 / 4) + BB * (z - 1 / 2)\ wneuper@59536: eq_a = Take eq; \ \([5], Res), 3 = AA * (1 / 2 - -1 / 4) + BB * (1 / 2 - 1 / 2)\ wneuper@59536: eq_a = Substitute [zzz = z1] eq; \ \([6], Res), 3 = 3 * AA / 4\ walther@59635: eq_a = (Rewrite_Set ''norm_Rational'') eq_a; wneuper@59536: \ \----- ([7], Pbl), solve (3 = 3 * AA / 4, AA)\ wneuper@59536: \ \([7], Res), [AA = 4]\ wneuper@59592: sol_a = SubProblem (''Isac_Knowledge'', [''univariate'',''equation''], [''no_met'']) wneuper@59536: [BOOL eq_a, REAL (AA::real)] ; \ \a = 4\ wneuper@59536: a = rhs (NTH 1 sol_a); \ \([8], Frm), 3 = AA * (z - -1 / 4) + BB * (z - 1 / 2)\ wneuper@59536: eq_b = Take eq; \ \([8], Res), 3 = AA * (-1 / 4 - -1 / 4) + BB * (-1 / 4 - 1 / 2)\ wneuper@59536: eq_b = Substitute [zzz = z2] eq_b; \ \([9], Res), 3 = -3 * BB / 4\ walther@59635: eq_b = (Rewrite_Set ''norm_Rational'') eq_b; \ \([10], Pbl), solve (3 = -3 * BB / 4, BB)\ walther@59635: sol_b = SubProblem (''Isac_Knowledge'', \ \([10], Res), [BB = -4]\ wneuper@59504: [''univariate'',''equation''], [''no_met'']) wneuper@59536: [BOOL eq_b, REAL (BB::real)]; \ \b = -4\ wneuper@59536: b = rhs (NTH 1 sol_b); \ \eqr = AA / (z - 1 / 2) + BB / (z - -1 / 4)\ wneuper@59504: pbz = Take eqr; \ \([11], Res), 4 / (z - 1 / 2) + -4 / (z - -1 / 4)\ wneuper@59512: pbz = Substitute [AA = a, BB = b] pbz \ \([], Res), 4 / (z - 1 / 2) + -4 / (z - -1 / 4)\ wneuper@59504: in pbz) " wenzelm@60303: wenzelm@60303: method met_partial_fraction : "simplification/of_rationals/to_partial_fraction" = wenzelm@60303: \{rew_ord'="tless_true", rls'= Rule_Set.empty, calc = [], srls = srls_partial_fraction, prls = Rule_Set.empty, wenzelm@60303: crls = Rule_Set.empty, errpats = [], nrls = Rule_Set.empty}\ wenzelm@60303: (*f_f = 3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z)), zzz: z*) wenzelm@60303: (*([], Frm), Problem (Partial_Fractions, [partial_fraction, rational, simplification])*) wenzelm@60303: Program: partial_fraction.simps wenzelm@60303: Given: "functionTerm t_t" "solveFor v_v" wenzelm@60303: (* Where: "((get_numerator t_t) has_degree_in v_v) < wenzelm@60303: ((get_denominator t_t) has_degree_in v_v)" TODO *) wenzelm@60303: Find: "decomposedFunction p_p'''" wenzelm@60303: wneuper@59472: ML \ neuper@42376: (* neuper@42376: val fmz = neuper@42376: ["functionTerm (3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z))))", neuper@42376: "solveFor z", "functionTerm p_p"]; neuper@42376: val (dI',pI',mI') = neuper@42376: ("Partial_Fractions", neuper@42376: ["partial_fraction", "rational", "simplification"], walther@59997: ["simplification", "of_rationals", "to_partial_fraction"]); neuper@42376: val (p,_,f,nxt,_,pt) = CalcTreeTEST [(fmz, (dI',pI',mI'))]; neuper@42376: *) walther@60278: \ML \ walther@60278: \ ML \ wneuper@59472: \ neuper@42289: end