src/Tools/isac/Knowledge/Inverse_Z_Transform.thy
author Walther Neuper <wneuper@ist.tugraz.at>
Wed, 28 Nov 2018 11:46:00 +0100
changeset 59473 28b67cae58c3
parent 59472 3e904f8ec16c
child 59489 cfcbcac0bae8
permissions -rw-r--r--
funpack: separate programs in prep. for partial_function
     1 (* Title:  Test_Z_Transform
     2    Author: Jan Rocnik
     3    (c) copyright due to lincense terms.
     4 12345678901234567890123456789012345678901234567890123456789012345678901234567890
     5         10        20        30        40        50        60        70        80
     6 *)
     7 
     8 theory Inverse_Z_Transform imports PolyEq DiffApp Partial_Fractions begin
     9 
    10 axiomatization where 
    11   rule1: "1 = \<delta>[n]" and
    12   rule2: "|| z || > 1 ==> z / (z - 1) = u [n]" and
    13   rule3: "|| z || < 1 ==> z / (z - 1) = -u [-n - 1]" and 
    14   rule4: "c * (z / (z - \<alpha>)) = c * \<alpha>^^^n * u [n]" and
    15   rule5: "|| z || < || \<alpha> || ==> z / (z - \<alpha>) = -(\<alpha>^^^n) * u [-n - 1]" and
    16   rule6: "|| z || > 1 ==> z/(z - 1)^^^2 = n * u [n]" (*and
    17   rule42: "(a * (z/(z-b)) + c * (z/(z-d))) = (a * b^^^n * u [n] + c * d^^^n * u [n])"*)
    18 
    19 axiomatization where
    20   ruleZY: "(X z = a / b) = (X' z = a / (z * b))" and
    21   ruleYZ: "(a/b + c/d) = (a*(z/b) + c*(z/d))" 
    22 
    23 subsection\<open>Define the Field Descriptions for the specification\<close>
    24 consts
    25   filterExpression  :: "bool => una"
    26   stepResponse      :: "bool => una"
    27 
    28 
    29 ML \<open>
    30 val inverse_z = prep_rls'(
    31   Rule.Rls {id = "inverse_z", preconds = [], rew_ord = ("dummy_ord",Rule.dummy_ord), 
    32 	  erls = Rule.Erls, srls = Rule.Erls, calc = [], errpatts = [],
    33 	  rules = 
    34 	   [
    35     Rule.Thm ("rule4", @{thm rule4})
    36 	   ], 
    37 	 scr = Rule.EmptyScr});
    38 \<close>
    39 
    40 
    41 text \<open>store the rule set for math engine\<close>
    42 
    43 setup \<open>KEStore_Elems.add_rlss [("inverse_z", (Context.theory_name @{theory}, inverse_z))]\<close>
    44 
    45 subsection\<open>Define the Specification\<close>
    46 ML \<open>
    47 val thy = @{theory};
    48 \<close>
    49 setup \<open>KEStore_Elems.add_pbts
    50   [(Specify.prep_pbt thy "pbl_SP" [] Celem.e_pblID (["SignalProcessing"], [], Rule.e_rls, NONE, [])),
    51     (Specify.prep_pbt thy "pbl_SP_Ztrans" [] Celem.e_pblID
    52       (["Z_Transform","SignalProcessing"], [], Rule.e_rls, NONE, [])),
    53     (Specify.prep_pbt thy "pbl_SP_Ztrans_inv" [] Celem.e_pblID
    54       (["Inverse", "Z_Transform", "SignalProcessing"],
    55         (*^ capital letter breaks coding standard
    56           because "inverse" = Const ("Rings.inverse_class.inverse", ..*)
    57         [("#Given" ,["filterExpression (X_eq::bool)"]),
    58           ("#Find"  ,["stepResponse (n_eq::bool)"])],
    59         Rule.append_rls "e_rls" Rule.e_rls [(*for preds in where_*)], NONE, 
    60         [["SignalProcessing","Z_Transform","Inverse"]])),
    61     (Specify.prep_pbt thy "pbl_SP_Ztrans_inv" [] Celem.e_pblID
    62       (["Inverse", "Z_Transform", "SignalProcessing"],
    63         [("#Given" ,["filterExpression X_eq"]),
    64           ("#Find"  ,["stepResponse n_eq"])],
    65         Rule.append_rls "e_rls" Rule.e_rls [(*for preds in where_*)], NONE, 
    66         [["SignalProcessing","Z_Transform","Inverse"]]))]\<close>
    67 
    68 subsection \<open>Define Name and Signature for the Method\<close>
    69 consts
    70   InverseZTransform :: "[bool, bool] => bool"
    71     ("((Script InverseZTransform (_ =))// (_))" 9)
    72 
    73 subsection \<open>Setup Parent Nodes in Hierarchy of Method\<close>
    74 ML \<open>val thy = @{theory}; (*latest version of thy required*)\<close>
    75 setup \<open>KEStore_Elems.add_mets
    76     [Specify.prep_met thy "met_SP" [] Celem.e_metID
    77       (["SignalProcessing"], [],
    78         {rew_ord'="tless_true", rls'= Rule.e_rls, calc = [], srls = Rule.e_rls, prls = Rule.e_rls, crls = Rule.e_rls,
    79           errpats = [], nrls = Rule.e_rls}, "empty_script"),
    80     Specify.prep_met thy "met_SP_Ztrans" [] Celem.e_metID
    81       (["SignalProcessing", "Z_Transform"], [],
    82         {rew_ord'="tless_true", rls'= Rule.e_rls, calc = [], srls = Rule.e_rls, prls = Rule.e_rls, crls = Rule.e_rls,
    83           errpats = [], nrls = Rule.e_rls}, "empty_script")]
    84 \<close>
    85 setup \<open>KEStore_Elems.add_mets
    86     [Specify.prep_met thy "met_SP_Ztrans_inv" [] Celem.e_metID
    87       (["SignalProcessing", "Z_Transform", "Inverse"], 
    88         [("#Given" ,["filterExpression (X_eq::bool)"]),
    89           ("#Find"  ,["stepResponse (n_eq::bool)"])],
    90         {rew_ord'="tless_true", rls'= Rule.e_rls, calc = [], srls = Rule.e_rls, prls = Rule.e_rls, crls = Rule.e_rls,
    91           errpats = [], nrls = Rule.e_rls},
    92         "Script InverseZTransform (X_eq::bool) =" ^ (*(1/z) instead of z ^^^ -1*)
    93           " (let X = Take X_eq;" ^
    94           "      X' = Rewrite ruleZY False X;" ^ (*z * denominator*)
    95           "      X' = (Rewrite_Set norm_Rational False) X';" ^ (*simplify*)
    96           "      funterm = Take (rhs X');" ^ (*drop X' z = for equation solving*)
    97           "      denom = (Rewrite_Set partial_fraction False) funterm;" ^ (*get_denominator*)
    98           "      equ = (denom = (0::real));" ^
    99           "      fun_arg = Take (lhs X');" ^
   100           "      arg = (Rewrite_Set partial_fraction False) X';" ^ (*get_argument TODO*)
   101           "      (L_L::bool list) =                                    " ^
   102           "            (SubProblem (Test',                            " ^
   103           "                         [LINEAR,univariate,equation,test]," ^
   104           "                         [Test,solve_linear])              " ^
   105           "                        [BOOL equ, REAL z])              " ^
   106           "  in X)")]
   107 \<close>
   108 setup \<open>KEStore_Elems.add_mets
   109     [Specify.prep_met thy "met_SP_Ztrans_inv" [] Celem.e_metID
   110       (["SignalProcessing", "Z_Transform", "Inverse"], 
   111         [("#Given" ,["filterExpression X_eq"]),
   112           ("#Find"  ,["stepResponse n_eq"])],
   113         {rew_ord'="tless_true", rls'= Rule.e_rls, calc = [], srls = srls_partial_fraction, prls = Rule.e_rls,
   114           crls = Rule.e_rls, errpats = [], nrls = Rule.e_rls},
   115         "Script InverseZTransform (X_eq::bool) =                        "^
   116            (*(1/z) instead of z ^^^ -1*)
   117            "(let X = Take X_eq;                                            "^
   118            "      X' = Rewrite ruleZY False X;                             "^
   119            (*z * denominator*)
   120            "      (num_orig::real) = get_numerator (rhs X');               "^
   121            "      X' = (Rewrite_Set norm_Rational False) X';               "^
   122            (*simplify*)
   123            "      (X'_z::real) = lhs X';                                   "^
   124            "      (zzz::real) = argument_in X'_z;                          "^
   125            "      (funterm::real) = rhs X';                                "^
   126            (*drop X' z = for equation solving*)
   127            "      (denom::real) = get_denominator funterm;                 "^
   128            (*get_denominator*)
   129            "      (num::real) = get_numerator funterm;                     "^
   130            (*get_numerator*)
   131            "      (equ::bool) = (denom = (0::real));                       "^
   132            "      (L_L::bool list) = (SubProblem (PolyEq',                 "^
   133            "         [abcFormula,degree_2,polynomial,univariate,equation], "^
   134            "         [no_met])                                             "^
   135            "         [BOOL equ, REAL zzz]);                                "^
   136            "      (facs::real) = factors_from_solution L_L;                "^
   137            "      (eql::real) = Take (num_orig / facs);                    "^ 
   138       
   139            "      (eqr::real) = (Try (Rewrite_Set ansatz_rls False)) eql;  "^
   140       
   141            "      (eq::bool) = Take (eql = eqr);                           "^
   142            (*Maybe possible to use HOLogic.mk_eq ??*)
   143            "      eq = (Try (Rewrite_Set equival_trans False)) eq;         "^ 
   144       
   145            "      eq = drop_questionmarks eq;                              "^
   146            "      (z1::real) = (rhs (NTH 1 L_L));                          "^
   147            (* 
   148             * prepare equation for a - eq_a
   149             * therefor substitute z with solution 1 - z1
   150             *)
   151            "      (z2::real) = (rhs (NTH 2 L_L));                          "^
   152        
   153            "      (eq_a::bool) = Take eq;                                  "^
   154            "      eq_a = (Substitute [zzz=z1]) eq;                         "^
   155            "      eq_a = (Rewrite_Set norm_Rational False) eq_a;           "^
   156            "      (sol_a::bool list) =                                     "^
   157            "                 (SubProblem (Isac',                           "^
   158            "                              [univariate,equation],[no_met])  "^
   159            "                              [BOOL eq_a, REAL (A::real)]);    "^
   160            "      (a::real) = (rhs(NTH 1 sol_a));                          "^
   161       
   162            "      (eq_b::bool) = Take eq;                                  "^
   163            "      eq_b =  (Substitute [zzz=z2]) eq_b;                      "^
   164            "      eq_b = (Rewrite_Set norm_Rational False) eq_b;           "^
   165            "      (sol_b::bool list) =                                     "^
   166            "                 (SubProblem (Isac',                           "^
   167            "                              [univariate,equation],[no_met])  "^
   168            "                              [BOOL eq_b, REAL (B::real)]);    "^
   169            "      (b::real) = (rhs(NTH 1 sol_b));                          "^
   170       
   171            "      eqr = drop_questionmarks eqr;                            "^
   172            "      (pbz::real) = Take eqr;                                  "^
   173            "      pbz = ((Substitute [A=a, B=b]) pbz);                     "^
   174       
   175            "      pbz = Rewrite ruleYZ False pbz;                          "^
   176            "      pbz = drop_questionmarks pbz;                            "^
   177       
   178            "      (X_z::bool) = Take (X_z = pbz);                          "^
   179            "      (n_eq::bool) = (Rewrite_Set inverse_z False) X_z;        "^
   180            "      n_eq = drop_questionmarks n_eq                           "^
   181            "in n_eq)")]
   182 \<close>
   183 setup \<open>KEStore_Elems.add_mets
   184     [Specify.prep_met thy "met_SP_Ztrans_inv_sub" [] Celem.e_metID
   185       (["SignalProcessing", "Z_Transform", "Inverse_sub"], 
   186         [("#Given" ,["filterExpression X_eq"]),
   187           ("#Find"  ,["stepResponse n_eq"])],
   188         {rew_ord'="tless_true", rls'= Rule.e_rls, calc = [],
   189           srls = Rule.Rls {id="srls_partial_fraction", 
   190               preconds = [], rew_ord = ("termlessI",termlessI),
   191               erls = Rule.append_rls "erls_in_srls_partial_fraction" Rule.e_rls
   192                   [(*for asm in NTH_CONS ...*)
   193                     Rule.Calc ("Orderings.ord_class.less",eval_equ "#less_"),
   194                     (*2nd NTH_CONS pushes n+-1 into asms*)
   195                     Rule.Calc("Groups.plus_class.plus", eval_binop "#add_")], 
   196               srls = Rule.Erls, calc = [], errpatts = [],
   197               rules = [Rule.Thm ("NTH_CONS", @{thm NTH_CONS}),
   198                   Rule.Calc ("Groups.plus_class.plus", eval_binop "#add_"),
   199                   Rule.Thm ("NTH_NIL", @{thm NTH_NIL}),
   200                   Rule.Calc ("Tools.lhs", eval_lhs "eval_lhs_"),
   201                   Rule.Calc ("Tools.rhs", eval_rhs"eval_rhs_"),
   202                   Rule.Calc ("Atools.argument'_in", eval_argument_in "Atools.argument'_in"),
   203                   Rule.Calc ("Rational.get_denominator", eval_get_denominator "#get_denominator"),
   204                   Rule.Calc ("Rational.get_numerator", eval_get_numerator "#get_numerator"),
   205                   Rule.Calc ("Partial_Fractions.factors_from_solution",
   206                     eval_factors_from_solution "#factors_from_solution"),
   207                   Rule.Calc("Partial_Fractions.drop_questionmarks", eval_drop_questionmarks "#drop_?")],
   208               scr = Rule.EmptyScr},
   209           prls = Rule.e_rls, crls = Rule.e_rls, errpats = [], nrls = norm_Rational},
   210         (*([], Frm), Problem (Isac, [Inverse, Z_Transform, SignalProcessing])*)
   211         "Script InverseZTransform (X_eq::bool) =            "^
   212           (*([1], Frm), X z = 3 / (z - 1 / 4 + -1 / 8 * (1 / z))*)
   213           "(let X = Take X_eq;                                "^
   214           (*([1], Res), ?X' z = 3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z)))*)
   215           "  X' = Rewrite ruleZY False X;                     "^
   216           (*            ?X' z*)
   217           "  (X'_z::real) = lhs X';                           "^
   218           (*            z *)
   219           "  (zzz::real) = argument_in X'_z;                  "^
   220           (*            3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z)))*)
   221           "  (funterm::real) = rhs X';                        "^
   222 
   223           "  (pbz::real) = (SubProblem (Isac',                "^
   224           "    [partial_fraction,rational,simplification],    "^
   225           "    [simplification,of_rationals,to_partial_fraction]) "^
   226           (*([2], Res), 4 / (z - 1 / 2) + -4 / (z - -1 / 4)*)
   227           "    [REAL funterm, REAL zzz]);                     "^
   228 
   229           (*([3], Frm), ?X' z = 4 / (z - 1 / 2) + -4 / (z - -1 / 4)*)
   230           "  (pbz_eq::bool) = Take (X'_z = pbz);              "^
   231           (*([3], Res), ?X' z = 4 * (?z / (z - 1 / 2)) + -4 * (?z / (z - -1 / 4))*)
   232           "  pbz_eq = Rewrite ruleYZ False pbz_eq;            "^
   233           (*               4 * (z / (z - 1 / 2)) + -4 * (z / (z - -1 / 4))*)
   234           "  pbz_eq = drop_questionmarks pbz_eq;              "^
   235           (*([4], Frm), X_z = 4 * (z / (z - 1 / 2)) + -4 * (z / (z - -1 / 4))*)
   236           "  (X_zeq::bool) = Take (X_z = rhs pbz_eq);         "^
   237           (*([4], Res), X_z = 4 * (1 / 2) ^^^ ?n * ?u [?n] + -4 * (-1 / 4) ^^^ ?n * ?u [?n]*)
   238           "  n_eq = (Rewrite_Set inverse_z False) X_zeq;      "^
   239           (*            X_z = 4 * (1 / 2) ^^^ n * u [n] + -4 * (-1 / 4) ^^^ n * u [n]*)
   240           "  n_eq = drop_questionmarks n_eq                   "^
   241           (*([], Res), X_z = 4 * (1 / 2) ^^^ n * u [n] + -4 * (-1 / 4) ^^^ n * u [n]*)
   242           "in n_eq)")]
   243 \<close>
   244 
   245 end
   246