src/Tools/isac/Knowledge/Inverse_Z_Transform.thy
author wneuper <Walther.Neuper@jku.at>
Thu, 04 Aug 2022 12:48:37 +0200
changeset 60509 2e0b7ca391dc
parent 60458 af7735fd252f
child 60515 03e19793d81e
permissions -rw-r--r--
polish naming in Rewrite_Order
     1 (* Title:  Inverse_Z_Transform
     2    Author: Jan Rocnik
     3    (c) copyright due to lincense terms.
     4 *)
     5 
     6 theory Inverse_Z_Transform imports PolyEq Diff_App Partial_Fractions begin
     7 
     8 axiomatization where       \<comment> \<open>TODO: new variables on the rhs enforce replacement by substitution\<close>
     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> \<up> n * u [n]" and
    13   rule5: "|| z || < || \<alpha> || ==> z / (z - \<alpha>) = -(\<alpha> \<up> n) * u [-n - 1]" and
    14   rule6: "|| z || > 1 ==> z/(z - 1) \<up> 2 = n * u [n]" (*and
    15   rule42: "(a * (z/(z-b)) + c * (z/(z-d))) = (a * b \<up> n * u [n] + c * d \<up> 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"    \<comment> \<open>TODO: unused, "u [n]" is introduced by rule1..6 above\<close>
    27 
    28 ML \<open>
    29 val inverse_z = prep_rls'(
    30   Rule_Def.Repeat {id = "inverse_z", preconds = [], rew_ord = ("dummy_ord",Rewrite_Ord.function_empty), 
    31 	  erls = Rule_Set.Empty, srls = Rule_Set.Empty, calc = [], errpatts = [],
    32 	  rules = 
    33 	   [
    34     \<^rule_thm>\<open>rule4\<close>
    35 	   ], 
    36 	 scr = Rule.Empty_Prog});
    37 \<close>
    38 
    39 
    40 text \<open>store the rule set for math engine\<close>
    41 
    42 rule_set_knowledge inverse_z = inverse_z
    43 
    44 subsection\<open>Define the Specification\<close>
    45 
    46 problem pbl_SP : "SignalProcessing" = \<open>Rule_Set.empty\<close>
    47 
    48 problem pbl_SP_Ztrans : "Z_Transform/SignalProcessing" = \<open>Rule_Set.empty\<close>
    49 
    50 problem pbl_SP_Ztrans_inv : "Inverse/Z_Transform/SignalProcessing" =
    51   \<open>Rule_Set.append_rules "empty" Rule_Set.empty [(*for preds in where_*)]\<close>
    52   Method_Ref: "SignalProcessing/Z_Transform/Inverse"
    53   Given: "filterExpression X_eq"
    54   Find: "stepResponse n_eq" \<comment> \<open>TODO: unused, "u [n]" is introduced by rule1..6\<close>
    55 
    56 
    57 subsection \<open>Setup Parent Nodes in Hierarchy of MethodC\<close>
    58 
    59 method met_SP : "SignalProcessing" =
    60   \<open>{rew_ord'="tless_true", rls'= Rule_Set.empty, calc = [], srls = Rule_Set.empty, prls = Rule_Set.empty, crls = Rule_Set.empty,
    61     errpats = [], nrls = Rule_Set.empty}\<close>
    62 
    63 method met_SP_Ztrans : "SignalProcessing/Z_Transform" =
    64   \<open>{rew_ord'="tless_true", rls'= Rule_Set.empty, calc = [], srls = Rule_Set.empty, prls = Rule_Set.empty, crls = Rule_Set.empty,
    65     errpats = [], nrls = Rule_Set.empty}\<close>
    66 
    67 partial_function (tailrec) inverse_ztransform :: "bool \<Rightarrow> real \<Rightarrow> bool"
    68   where
    69 "inverse_ztransform X_eq X_z = (
    70   let
    71     X = Take X_eq;
    72     X' = Rewrite ''ruleZY''  X;                                              \<comment> \<open>z * denominator\<close>
    73     X' = (Rewrite_Set ''norm_Rational'' ) X';                                       \<comment> \<open>simplify\<close>
    74     funterm = Take (rhs X');                                \<comment> \<open>drop X' z = for equation solving\<close>
    75     denom = (Rewrite_Set ''partial_fraction'' ) funterm;                     \<comment> \<open>get_denominator\<close>
    76     equ = (denom = (0::real));
    77     fun_arg = Take (lhs X');
    78     arg = (Rewrite_Set ''partial_fraction'' ) X';                          \<comment> \<open>get_argument TODO\<close>
    79     (L_L::bool list) = \<comment> \<open>'bool list' inhibits:
    80                 WARNING: Additional type variable(s) in specification of inverse_ztransform: 'a\<close>
    81       SubProblem (''Test'', [''LINEAR'',''univariate'',''equation'',''test''], [''Test'',''solve_linear''])
    82         [BOOL equ, REAL X_z]
    83   in X) "
    84 
    85 method met_SP_Ztrans_inv : "SignalProcessing/Z_Transform/Inverse" =
    86   \<open>{rew_ord'="tless_true", rls'= Rule_Set.empty, calc = [], srls = Rule_Set.empty, prls = Rule_Set.empty, crls = Rule_Set.empty,
    87     errpats = [], nrls = Rule_Set.empty}\<close>
    88   Program: inverse_ztransform.simps
    89   Given: "filterExpression X_eq" "functionName X_z"
    90   Find: "stepResponse n_eq" \<comment> \<open>TODO: unused, "u [n]" is introduced by rule1..6\<close>
    91 
    92 partial_function (tailrec) inverse_ztransform2 :: "bool \<Rightarrow> real \<Rightarrow> bool"
    93   where
    94 "inverse_ztransform2 X_eq X_z = (
    95   let
    96     X = Take X_eq;
    97     X' = Rewrite ''ruleZY''  X;
    98     X_z = lhs X';
    99     zzz = argument_in X_z;
   100     funterm = rhs X';
   101     pbz = SubProblem (''Isac_Knowledge'',
   102         [''partial_fraction'',''rational'',''simplification''],
   103         [''simplification'',''of_rationals'',''to_partial_fraction''])
   104       [REAL funterm, REAL zzz];
   105     pbz_eq = Take (X_z = pbz);
   106     pbz_eq = Rewrite ''ruleYZ''  pbz_eq;
   107     X_zeq = Take (X_z = rhs pbz_eq);
   108     n_eq = (Rewrite_Set ''inverse_z'' ) X_zeq
   109   in n_eq)"
   110 
   111 method met_SP_Ztrans_inv_sub : "SignalProcessing/Z_Transform/Inverse_sub" =
   112   \<open>{rew_ord'="tless_true", rls'= Rule_Set.empty, calc = [],
   113     srls = Rule_Def.Repeat {
   114       id = "srls_partial_fraction",  preconds = [], rew_ord = ("termlessI",termlessI),
   115       erls = Rule_Set.append_rules "erls_in_srls_partial_fraction" Rule_Set.empty [
   116         \<^rule_eval>\<open>less\<close> (Prog_Expr.eval_equ "#less_"), (* ...for asm in NTH_CONS*)
   117         (*2nd NTH_CONS pushes n+-1 into asms*)
   118         \<^rule_eval>\<open>plus\<close> (**)(eval_binop "#add_")], 
   119       srls = Rule_Set.Empty, calc = [], errpatts = [],
   120       rules = [\<^rule_thm>\<open>NTH_CONS\<close>,
   121         \<^rule_eval>\<open>plus\<close> (**)(eval_binop "#add_"),
   122         \<^rule_thm>\<open>NTH_NIL\<close>,
   123         \<^rule_eval>\<open>Prog_Expr.lhs\<close> (Prog_Expr.eval_lhs "eval_lhs_"),
   124         \<^rule_eval>\<open>Prog_Expr.rhs\<close> (Prog_Expr.eval_rhs"eval_rhs_"),
   125         \<^rule_eval>\<open>Prog_Expr.argument_in\<close> (Prog_Expr.eval_argument_in "Prog_Expr.argument_in"),
   126         \<^rule_eval>\<open>get_denominator\<close> (eval_get_denominator "#get_denominator"),
   127         \<^rule_eval>\<open>get_numerator\<close> (eval_get_numerator "#get_numerator"),
   128         \<^rule_eval>\<open>factors_from_solution\<close> (eval_factors_from_solution "#factors_from_solution")],
   129       scr = Rule.Empty_Prog},
   130     prls = Rule_Set.empty, crls = Rule_Set.empty, errpats = [], nrls = norm_Rational}\<close>
   131   Program: inverse_ztransform2.simps
   132   Given: "filterExpression X_eq" "functionName X_z"
   133   Find: "stepResponse n_eq" \<comment> \<open>TODO: unused, "u [n]" is introduced by rule1..6\<close>
   134 
   135 ML \<open>
   136 \<close> ML \<open>
   137 \<close> ML \<open>
   138 \<close>
   139 
   140 end
   141