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