src/Tools/isac/Knowledge/Inverse_Z_Transform.thy
author Walther Neuper <walther.neuper@jku.at>
Wed, 08 Apr 2020 12:32:51 +0200
changeset 59852 ea7e6679080e
parent 59851 4dd533681fef
child 59857 cbb3fae0381d
permissions -rw-r--r--
use new struct "Rule_Set" for renaming identifiers
     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 DiffApp 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>^^^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"    \<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",Rule.dummy_ord), 
    31 	  erls = Rule_Set.Empty, srls = Rule_Set.Empty, 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_Set.empty, NONE, [])),
    50     (Specify.prep_pbt thy "pbl_SP_Ztrans" [] Celem.e_pblID
    51       (["Z_Transform","SignalProcessing"], [], Rule_Set.empty, 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"])], \<comment> \<open>TODO: unused, "u [n]" is introduced by rule1..6\<close>
    56         Rule_Set.append_rules "empty" Rule_Set.empty [(*for preds in where_*)], NONE, 
    57         [["SignalProcessing","Z_Transform","Inverse"]]))]\<close>
    58 
    59 subsection \<open>Setup Parent Nodes in Hierarchy of Method\<close>
    60 ML \<open>val thy = @{theory}; (*latest version of thy required*)\<close>
    61 setup \<open>KEStore_Elems.add_mets
    62     [Specify.prep_met thy "met_SP" [] Celem.e_metID
    63       (["SignalProcessing"], [],
    64         {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}, @{thm refl}),
    66     Specify.prep_met thy "met_SP_Ztrans" [] Celem.e_metID
    67       (["SignalProcessing", "Z_Transform"], [],
    68         {rew_ord'="tless_true", rls'= Rule_Set.empty, calc = [], srls = Rule_Set.empty, prls = Rule_Set.empty, crls = Rule_Set.empty,
    69           errpats = [], nrls = Rule_Set.empty}, @{thm refl})]
    70 \<close>
    71 
    72 partial_function (tailrec) inverse_ztransform :: "bool \<Rightarrow> real \<Rightarrow> bool"
    73   where
    74 "inverse_ztransform X_eq X_z = (
    75   let
    76     X = Take X_eq;
    77     X' = Rewrite ''ruleZY''  X;                                              \<comment> \<open>z * denominator\<close>
    78     X' = (Rewrite_Set ''norm_Rational'' ) X';                                       \<comment> \<open>simplify\<close>
    79     funterm = Take (rhs X');                                \<comment> \<open>drop X' z = for equation solving\<close>
    80     denom = (Rewrite_Set ''partial_fraction'' ) funterm;                     \<comment> \<open>get_denominator\<close>
    81     equ = (denom = (0::real));
    82     fun_arg = Take (lhs X');
    83     arg = (Rewrite_Set ''partial_fraction'' ) X';                          \<comment> \<open>get_argument TODO\<close>
    84     (L_L::bool list) = \<comment> \<open>'bool list' inhibits:
    85                 WARNING: Additional type variable(s) in specification of inverse_ztransform: 'a\<close>
    86       SubProblem (''Test'', [''LINEAR'',''univariate'',''equation'',''test''], [''Test'',''solve_linear''])
    87         [BOOL equ, REAL X_z]
    88   in X) "
    89 setup \<open>KEStore_Elems.add_mets
    90     [Specify.prep_met thy "met_SP_Ztrans_inv" [] Celem.e_metID
    91       (["SignalProcessing", "Z_Transform", "Inverse"],
    92         [("#Given" ,["filterExpression X_eq", "functionName X_z"]),
    93           ("#Find"  ,["stepResponse n_eq"])], \<comment> \<open>TODO: unused, "u [n]" is introduced by rule1..6\<close>
    94         {rew_ord'="tless_true", rls'= Rule_Set.empty, calc = [], srls = Rule_Set.empty, prls = Rule_Set.empty, crls = Rule_Set.empty,
    95           errpats = [], nrls = Rule_Set.empty},
    96         @{thm inverse_ztransform.simps})]
    97 \<close>
    98 
    99 partial_function (tailrec) inverse_ztransform2 :: "bool \<Rightarrow> real \<Rightarrow> bool"
   100   where
   101 "inverse_ztransform2 X_eq X_z = (
   102   let
   103     X = Take X_eq;
   104     X' = Rewrite ''ruleZY''  X;
   105     X'_z = lhs X';
   106     zzz = argument_in X'_z;
   107     funterm = rhs X';
   108     pbz = SubProblem (''Isac_Knowledge'',
   109         [''partial_fraction'',''rational'',''simplification''],
   110         [''simplification'',''of_rationals'',''to_partial_fraction''])
   111       [REAL funterm, REAL zzz];
   112     pbz_eq = Take (X'_z = pbz);
   113     pbz_eq = Rewrite ''ruleYZ''  pbz_eq;
   114     X_zeq = Take (X_z = rhs pbz_eq);
   115     n_eq = (Rewrite_Set ''inverse_z'' ) X_zeq
   116   in n_eq)"
   117 setup \<open>KEStore_Elems.add_mets
   118     [Specify.prep_met thy "met_SP_Ztrans_inv_sub" [] Celem.e_metID
   119       (["SignalProcessing", "Z_Transform", "Inverse_sub"],
   120         [("#Given" ,["filterExpression X_eq", "functionName X_z"]),
   121           ("#Find"  ,["stepResponse n_eq"])], \<comment> \<open>TODO: unused, "u [n]" is introduced by rule1..6\<close>
   122         {rew_ord'="tless_true", rls'= Rule_Set.empty, calc = [],
   123           srls = Rule_Def.Repeat {id="srls_partial_fraction", 
   124               preconds = [], rew_ord = ("termlessI",termlessI),
   125               erls = Rule_Set.append_rules "erls_in_srls_partial_fraction" Rule_Set.empty
   126                   [(*for asm in NTH_CONS ...*)
   127                     Rule.Num_Calc ("Orderings.ord_class.less", Prog_Expr.eval_equ "#less_"),
   128                     (*2nd NTH_CONS pushes n+-1 into asms*)
   129                     Rule.Num_Calc("Groups.plus_class.plus", (**)eval_binop "#add_")], 
   130               srls = Rule_Set.Empty, calc = [], errpatts = [],
   131               rules = [Rule.Thm ("NTH_CONS", @{thm NTH_CONS}),
   132                   Rule.Num_Calc ("Groups.plus_class.plus", (**)eval_binop "#add_"),
   133                   Rule.Thm ("NTH_NIL", @{thm NTH_NIL}),
   134                   Rule.Num_Calc ("Prog_Expr.lhs", Prog_Expr.eval_lhs "eval_lhs_"),
   135                   Rule.Num_Calc ("Prog_Expr.rhs", Prog_Expr.eval_rhs"eval_rhs_"),
   136                   Rule.Num_Calc ("Prog_Expr.argument'_in", Prog_Expr.eval_argument_in "Prog_Expr.argument'_in"),
   137                   Rule.Num_Calc ("Rational.get_denominator", eval_get_denominator "#get_denominator"),
   138                   Rule.Num_Calc ("Rational.get_numerator", eval_get_numerator "#get_numerator"),
   139                   Rule.Num_Calc ("Partial_Fractions.factors_from_solution",
   140                     eval_factors_from_solution "#factors_from_solution")
   141                   ], scr = Rule.EmptyScr},
   142           prls = Rule_Set.empty, crls = Rule_Set.empty, errpats = [], nrls = norm_Rational},
   143         @{thm inverse_ztransform2.simps})]
   144 \<close>
   145 ML \<open>
   146 \<close> ML \<open>
   147 \<close>
   148 
   149 end
   150