src/Tools/isac/Knowledge/Inverse_Z_Transform.thy
author Walther Neuper <wneuper@ist.tugraz.at>
Wed, 29 May 2019 10:36:16 +0200
changeset 59545 4035ec339062
parent 59543 52a90a3c7881
child 59550 2e7631381921
permissions -rw-r--r--
[-Test_Isac] funpack: switch from Script to partial_function
     1 (* Title:  Test_Z_Transform
     2    Author: Jan Rocnik
     3    (c) copyright due to lincense terms.
     4 *)
     5 
     6 theory Inverse_Z_Transform imports PolyEq DiffApp Partial_Fractions begin
     7 
     8 axiomatization where 
     9   rule1: "1 = \<delta>[n]" and
    10   rule2: "|| z || > 1 ==> z / (z - 1) = u [n]" and
    11   rule3: "|| z || < 1 ==> z / (z - 1) = -u [-n - 1]" and 
    12   rule4: "c * (z / (z - \<alpha>)) = c * \<alpha>^^^n * u [n]" and
    13   rule5: "|| z || < || \<alpha> || ==> z / (z - \<alpha>) = -(\<alpha>^^^n) * u [-n - 1]" and
    14   rule6: "|| z || > 1 ==> z/(z - 1)^^^2 = n * u [n]" (*and
    15   rule42: "(a * (z/(z-b)) + c * (z/(z-d))) = (a * b^^^n * u [n] + c * d^^^n * u [n])"*)
    16 
    17 axiomatization where
    18 (*ruleZY: "(X z = a / b) = (d_d z X = a / (z * b))"         ..looks better, but types are flawed*)
    19   ruleZY: "(X z = a / b) = (X' z = a / (z * b))" and
    20   ruleYZ: "a / (z - b) + c / (z - d) = a * (z / (z - b)) + c * (z / (z - d))" and
    21   ruleYZa: "(a / b + c / d) = (a * (z / b) + c * (z / d))"        \<comment> \<open>that is what students learn\<close>
    22 
    23 subsection\<open>Define the Field Descriptions for the specification\<close>
    24 consts
    25   filterExpression  :: "bool => una"
    26   stepResponse      :: "bool => una"
    27 
    28 ML \<open>
    29 val inverse_z = prep_rls'(
    30   Rule.Rls {id = "inverse_z", preconds = [], rew_ord = ("dummy_ord",Rule.dummy_ord), 
    31 	  erls = Rule.Erls, srls = Rule.Erls, calc = [], errpatts = [],
    32 	  rules = 
    33 	   [
    34     Rule.Thm ("rule4", @{thm rule4})
    35 	   ], 
    36 	 scr = Rule.EmptyScr});
    37 \<close>
    38 
    39 
    40 text \<open>store the rule set for math engine\<close>
    41 
    42 setup \<open>KEStore_Elems.add_rlss [("inverse_z", (Context.theory_name @{theory}, inverse_z))]\<close>
    43 
    44 subsection\<open>Define the Specification\<close>
    45 ML \<open>
    46 val thy = @{theory};
    47 \<close>
    48 setup \<open>KEStore_Elems.add_pbts
    49   [(Specify.prep_pbt thy "pbl_SP" [] Celem.e_pblID (["SignalProcessing"], [], Rule.e_rls, NONE, [])),
    50     (Specify.prep_pbt thy "pbl_SP_Ztrans" [] Celem.e_pblID
    51       (["Z_Transform","SignalProcessing"], [], Rule.e_rls, NONE, [])),
    52     (Specify.prep_pbt thy "pbl_SP_Ztrans_inv" [] Celem.e_pblID
    53       (["Inverse", "Z_Transform", "SignalProcessing"],
    54         [("#Given" ,["filterExpression X_eq"]),
    55           ("#Find"  ,["stepResponse n_eq"])],
    56         Rule.append_rls "e_rls" Rule.e_rls [(*for preds in where_*)], NONE, 
    57         [["SignalProcessing","Z_Transform","Inverse"]])),
    58     (Specify.prep_pbt thy "pbl_SP_Ztrans_inv_sub" [] Celem.e_pblID
    59       (["Inverse_sub", "Z_Transform", "SignalProcessing"],
    60         [("#Given" ,["filterExpression X_eq"]),
    61           ("#Find"  ,["stepResponse n_eq"])],
    62         Rule.append_rls "e_rls" Rule.e_rls [(*for preds in where_*)], NONE, 
    63         [["SignalProcessing","Z_Transform","Inverse_sub"]]))]\<close>
    64 
    65 subsection \<open>Define Name and Signature for the Method\<close>
    66 consts
    67   InverseZTransform1 :: "[bool, bool] => bool"
    68     ("((Script InverseZTransform1 (_ =))// (_))" 9)
    69   InverseZTransform2 :: "[bool, real, bool] => bool"
    70     ("((Script InverseZTransform2 (_ _ =))// (_))" 9)
    71 
    72 subsection \<open>Setup Parent Nodes in Hierarchy of Method\<close>
    73 ML \<open>val thy = @{theory}; (*latest version of thy required*)\<close>
    74 setup \<open>KEStore_Elems.add_mets
    75     [Specify.prep_met thy "met_SP" [] Celem.e_metID
    76       (["SignalProcessing"], [],
    77         {rew_ord'="tless_true", rls'= Rule.e_rls, calc = [], srls = Rule.e_rls, prls = Rule.e_rls, crls = Rule.e_rls,
    78           errpats = [], nrls = Rule.e_rls}, @{thm refl}),
    79     Specify.prep_met thy "met_SP_Ztrans" [] Celem.e_metID
    80       (["SignalProcessing", "Z_Transform"], [],
    81         {rew_ord'="tless_true", rls'= Rule.e_rls, calc = [], srls = Rule.e_rls, prls = Rule.e_rls, crls = Rule.e_rls,
    82           errpats = [], nrls = Rule.e_rls}, @{thm refl})]
    83 \<close>
    84 
    85 (*WARNING: Additional type variable(s) in specification of "inverse_ztransform": 'a *)
    86 partial_function (tailrec) inverse_ztransform :: "bool \<Rightarrow> real \<Rightarrow> bool"
    87   where
    88 "inverse_ztransform X_eq z =                                           \<comment> \<open>(1/z) instead of z ^^^ -1\<close>                
    89  (let X = Take X_eq;                                                                
    90       X' = Rewrite ''ruleZY'' False X;                                         \<comment> \<open>z * denominator\<close>                          
    91       X' = (Rewrite_Set ''norm_Rational'' False) X';                                  \<comment> \<open>simplify\<close>                   
    92       funterm = Take (rhs X');                                \<comment> \<open>drop X' z = for equation solving\<close>                 
    93       denom = (Rewrite_Set ''partial_fraction'' False) funterm;                \<comment> \<open>get_denominator\<close> 
    94       equ = (denom = (0::real));                                                    
    95       fun_arg = Take (lhs X');                                                      
    96       arg = (Rewrite_Set ''partial_fraction'' False) X';                     \<comment> \<open>get_argument TODO\<close>      
    97       L_L = SubProblem (''Test'', [''LINEAR'',''univariate'',''equation'',''test''],         
    98                 [''Test'',''solve_linear'']) [BOOL equ, REAL z]               \<comment> \<open>PROG string\<close>
    99   in X) "
   100 setup \<open>KEStore_Elems.add_mets
   101     [Specify.prep_met thy "met_SP_Ztrans_inv" [] Celem.e_metID
   102       (["SignalProcessing", "Z_Transform", "Inverse"],
   103         [("#Given" ,["filterExpression (X_eq::bool)"]),
   104           ("#Find"  ,["stepResponse (n_eq::bool)"])],
   105         {rew_ord'="tless_true", rls'= Rule.e_rls, calc = [], srls = Rule.e_rls, prls = Rule.e_rls, crls = Rule.e_rls,
   106           errpats = [], nrls = Rule.e_rls},
   107         @{thm inverse_ztransform.simps}
   108 	    (*"Script InverseZTransform1 (X_eq::bool) =" ^ (*(1/z) instead of z ^^^ -1*)
   109           " (let X = Take X_eq;" ^
   110           "      X' = Rewrite ''ruleZY'' False X;" ^ (*z * denominator*)
   111           "      X' = (Rewrite_Set ''norm_Rational'' False) X';" ^ (*simplify*)
   112           "      funterm = Take (rhs X');" ^ (*drop X' z = for equation solving*)
   113           "      denom = (Rewrite_Set ''partial_fraction'' False) funterm;" ^ (*get_denominator*)
   114           "      equ = (denom = (0::real));" ^
   115           "      fun_arg = Take (lhs X');" ^
   116           "      arg = (Rewrite_Set ''partial_fraction'' False) X';" ^ (*get_argument TODO*)
   117           "      (L_L::bool list) =                                    " ^
   118           "            (SubProblem (''Test'',                            " ^
   119           "                         [''LINEAR'',''univariate'',''equation'',''test'']," ^
   120           "                         [''Test'',''solve_linear''])              " ^
   121           "                        [BOOL equ, REAL z])              " ^
   122           "  in X)"*))]
   123 \<close>
   124 
   125 partial_function (tailrec) inverse_ztransform2 :: "bool \<Rightarrow> real \<Rightarrow> bool"
   126   where
   127 "inverse_ztransform2 X_eq X_z =
   128   (let X = Take X_eq;
   129     X' = Rewrite ''ruleZY'' False X;
   130     X'_z = lhs X';
   131     zzz = argument_in X'_z;
   132     funterm = rhs X';
   133     pbz = SubProblem (''Isac'',
   134       [''partial_fraction'',''rational'',''simplification''],
   135       [''simplification'',''of_rationals'',''to_partial_fraction''])
   136       [REAL funterm, REAL zzz];
   137     pbz_eq = Take (X'_z = pbz);
   138     pbz_eq = Rewrite ''ruleYZ'' False pbz_eq;
   139     X_zeq = Take (X_z = rhs pbz_eq);
   140     n_eq = (Rewrite_Set ''inverse_z'' False) X_zeq
   141   in n_eq)"
   142 setup \<open>KEStore_Elems.add_mets
   143     [Specify.prep_met thy "met_SP_Ztrans_inv_sub" [] Celem.e_metID
   144       (["SignalProcessing", "Z_Transform", "Inverse_sub"],
   145         [("#Given" ,["filterExpression X_eq", "boundVariable X_z"]),
   146           ("#Find"  ,["stepResponse n_eq"])],
   147         {rew_ord'="tless_true", rls'= Rule.e_rls, calc = [],
   148           srls = Rule.Rls {id="srls_partial_fraction", 
   149               preconds = [], rew_ord = ("termlessI",termlessI),
   150               erls = Rule.append_rls "erls_in_srls_partial_fraction" Rule.e_rls
   151                   [(*for asm in NTH_CONS ...*)
   152                     Rule.Calc ("Orderings.ord_class.less",eval_equ "#less_"),
   153                     (*2nd NTH_CONS pushes n+-1 into asms*)
   154                     Rule.Calc("Groups.plus_class.plus", eval_binop "#add_")], 
   155               srls = Rule.Erls, calc = [], errpatts = [],
   156               rules = [Rule.Thm ("NTH_CONS", @{thm NTH_CONS}),
   157                   Rule.Calc ("Groups.plus_class.plus", eval_binop "#add_"),
   158                   Rule.Thm ("NTH_NIL", @{thm NTH_NIL}),
   159                   Rule.Calc ("Tools.lhs", Tools.eval_lhs "eval_lhs_"),
   160                   Rule.Calc ("Tools.rhs", Tools.eval_rhs"eval_rhs_"),
   161                   Rule.Calc ("Atools.argument'_in", eval_argument_in "Atools.argument'_in"),
   162                   Rule.Calc ("Rational.get_denominator", eval_get_denominator "#get_denominator"),
   163                   Rule.Calc ("Rational.get_numerator", eval_get_numerator "#get_numerator"),
   164                   Rule.Calc ("Partial_Fractions.factors_from_solution",
   165                     eval_factors_from_solution "#factors_from_solution")
   166                   ], scr = Rule.EmptyScr},
   167           prls = Rule.e_rls, crls = Rule.e_rls, errpats = [], nrls = norm_Rational},
   168         @{thm simplify.simps}
   169 	    (*" Script InverseZTransform2 (X_eq::bool) (X_z::real) =               "^ (*([], Frm), Problem (Isac, [Inverse, Z_Transform, SignalProcessing])*)
   170         " (let X = Take X_eq;                                                "^ (*([1], Frm), X z = 3 / (z - 1 / 4 + -1 / 8 * (1 / z))*)
   171         "   X' = Rewrite ''ruleZY'' False X;                                 "^ (*([1], Res), ?X' z = 3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z)))*)
   172         "   (X'_z::real) = lhs X';                                           "^ (*            ?X' z*)
   173         "   (zzz::real) = argument_in X'_z;                                  "^ (*            z *)
   174         "   (funterm::real) = rhs X';                                        "^ (*            3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z)))*)
   175         "   (pbz::real) = (SubProblem (''Isac'',                             "^ (*([2], Res), 4 / (z - 1 / 2) + -4 / (z - -1 / 4)*)
   176         "     [''partial_fraction'',''rational'',''simplification''],        "^
   177         "     [''simplification'',''of_rationals'',''to_partial_fraction'']) "^
   178         "     [REAL funterm, REAL zzz]);                                     "^
   179         "   (pbz_eq::bool) = Take (X'_z = pbz);                              "^ (*([3], Frm), ?X' z = 4 / (z - 1 / 2) + -4 / (z - -1 / 4)*)
   180         "   pbz_eq = Rewrite ''ruleYZ'' False pbz_eq;                        "^ (*([3], Res), ?X' z = 4 * (?z / (z - 1 / 2)) + -4 * (?z / (z - -1 / 4))*)
   181         "   (X_zeq::bool) = Take (X_z = rhs pbz_eq);                         "^ (*([4], Frm), X_z = 4 * (z / (z - 1 / 2)) + -4 * (z / (z - -1 / 4))*)
   182         "   n_eq = (Rewrite_Set ''inverse_z'' False) X_zeq                   "^ (*([4], Res), X_z = 4 * (1 / 2) ^^^ ?n * ?u [?n] + -4 * (-1 / 4) ^^^ ?n * ?u [?n]*)
   183         " in n_eq)                                                           "*))](*            X_z = 4 * (1 / 2) ^^^ n * u [n] + -4 * (-1 / 4) ^^^ n * u [n]*)
   184 \<close>
   185 
   186 end
   187