src/Tools/isac/Knowledge/Inverse_Z_Transform.thy
author Walther Neuper <wneuper@ist.tugraz.at>
Thu, 28 Feb 2019 12:14:32 +0100
changeset 59505 a1f223658994
parent 59504 546bd1b027cc
child 59512 e504168e7b01
permissions -rw-r--r--
funpack: outcomment partial_function, further preps required
     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 (*ok
    86 partial_function (tailrec) inverse_ztransform :: "bool \<Rightarrow> bool"
    87   where
    88 "inverse_ztransform X_eq =                                           \<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, STRING ''z'']              \<comment> \<open>PROG string\<close>
    99   in X) "
   100 *)
   101 setup \<open>KEStore_Elems.add_mets
   102     [Specify.prep_met thy "met_SP_Ztrans_inv" [] Celem.e_metID
   103       (["SignalProcessing", "Z_Transform", "Inverse"], 
   104         [("#Given" ,["filterExpression (X_eq::bool)"]),
   105           ("#Find"  ,["stepResponse (n_eq::bool)"])],
   106         {rew_ord'="tless_true", rls'= Rule.e_rls, calc = [], srls = Rule.e_rls, prls = Rule.e_rls, crls = Rule.e_rls,
   107           errpats = [], nrls = Rule.e_rls},
   108         "Script InverseZTransform (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 Type unification failed: Clash of types "bool" and "_ itself"
   126 Type error in application: incompatible operand type
   127 Operator:  Let (Take X_eq) :: (??'a itself \<Rightarrow> ??'b) \<Rightarrow> ??'b
   128 Operand:
   129   \<lambda>X. let X' = Rewrite ''ruleZY'' ...
   130 :
   131 :partial_function (tailrec) inverse_ztransform2 :: "bool \<Rightarrow> bool"
   132   where
   133 "inverse_ztransform X_eq =                                          \<comment> \<open>(1/z) instead of z ^^^ -1\<close>
   134   (let X = Take X_eq;
   135       X' = Rewrite ''ruleZY'' False X;                                        \<comment> \<open>z * denominator\<close>
   136       (num_orig::real) = get_numerator (rhs X');                                  
   137       X' = (Rewrite_Set ''norm_Rational'' False) X';                                 \<comment> \<open>simplify\<close>
   138       (X'_z::real) = lhs X';                                                      
   139       (zzz::real) = argument_in X'_z;                                             
   140       (funterm::real) = rhs X';                              \<comment> \<open>drop X' z = for equation solving\<close>
   141       (denom::real) = get_denominator funterm;                                \<comment> \<open>get_denominator\<close>
   142       (num::real) = get_numerator funterm;                                      \<comment> \<open>get_numerator\<close>
   143       (equ::bool) = (denom = (0::real));                                          
   144       (L_L::bool list) = (SubProblem (''PolyEq'',                                 
   145          [''abcFormula'',''degree_2'',''polynomial'',''univariate'',''equation''],
   146          [''no_met''])                                                            
   147          [BOOL equ, REAL zzz]);                                                   
   148       (facs::real) = factors_from_solution L_L;                                   
   149       (eql::real) = Take (num_orig / facs); \<comment> \<open>---\<close>
   150       (eqr::real) = (Try (Rewrite_Set ''ansatz_rls'' False)) eql; \<comment> \<open>---\<close>
   151       (eq::bool) = Take (eql = eqr);                    \<comment> \<open>Maybe possible to use HOLogic.mk_eq ??\<close>
   152       eq = (Try (Rewrite_Set ''equival_trans'' False)) eq; \<comment> \<open>---\<close>
   153       eq = drop_questionmarks eq;                                                 
   154       (z1::real) = (rhs (NTH 1 L_L)); \<comment> \<open>prepare equation for a - eq_a therefor substitute z with solution 1 - z1\<close>
   155       (z2::real) = (rhs (NTH 2 L_L)); \<comment> \<open>---\<close>
   156       (eq_a::bool) = Take eq;                                                     
   157       eq_a = (Substitute [zzz=z1]) eq;                                            
   158       eq_a = (Rewrite_Set ''norm_Rational'' False) eq_a;                          
   159       (sol_a::bool list) =                                                        
   160                  (SubProblem (''Isac'',                                           
   161                               [''univariate'',''equation''],[''no_met''])         
   162                               [BOOL eq_a, REAL (A::real)]);                       
   163       (a::real) = (rhs(NTH 1 sol_a)); \<comment> \<open>---\<close>
   164       (eq_b::bool) = Take eq;                                                     
   165       eq_b =  (Substitute [zzz=z2]) eq_b;                                         
   166       eq_b = (Rewrite_Set ''norm_Rational'' False) eq_b;                          
   167       (sol_b::bool list) =                                                        
   168                  (SubProblem (''Isac'',                                           
   169                               [''univariate'',''equation''],[''no_met''])         
   170                               [BOOL eq_b, REAL (B::real)]);                       
   171       (b::real) = (rhs(NTH 1 sol_b)); \<comment> \<open>---\<close>
   172       eqr = drop_questionmarks eqr;                                               
   173       (pbz::real) = Take eqr;                                                     
   174       pbz = ((Substitute [A=a, B=b]) pbz); \<comment> \<open>---\<close>
   175       pbz = Rewrite ''ruleYZ'' False pbz;                                         
   176       pbz = drop_questionmarks pbz; \<comment> \<open>---\<close>
   177       (X_z::bool) = Take (X_z = pbz);                                             
   178       (n_eq::bool) = (Rewrite_Set ''inverse_z'' False) X_z;                       
   179       n_eq = drop_questionmarks n_eq                                              
   180 in n_eq)"
   181 *)
   182 setup \<open>KEStore_Elems.add_mets
   183     [Specify.prep_met thy "met_SP_Ztrans_inv" [] Celem.e_metID
   184       (["SignalProcessing", "Z_Transform", "Inverse"], 
   185         [("#Given" ,["filterExpression X_eq"]),
   186           ("#Find"  ,["stepResponse n_eq"])],
   187         {rew_ord'="tless_true", rls'= Rule.e_rls, calc = [], srls = srls_partial_fraction, prls = Rule.e_rls,
   188           crls = Rule.e_rls, errpats = [], nrls = Rule.e_rls},
   189         "Script InverseZTransform (X_eq::bool) =                        "^
   190            (*(1/z) instead of z ^^^ -1*)
   191            "(let X = Take X_eq;                                            "^
   192            "      X' = Rewrite ''ruleZY'' False X;                             "^
   193            (*z * denominator*)
   194            "      (num_orig::real) = get_numerator (rhs X');               "^
   195            "      X' = (Rewrite_Set ''norm_Rational'' False) X';               "^
   196            (*simplify*)
   197            "      (X'_z::real) = lhs X';                                   "^
   198            "      (zzz::real) = argument_in X'_z;                          "^
   199            "      (funterm::real) = rhs X';                                "^
   200            (*drop X' z = for equation solving*)
   201            "      (denom::real) = get_denominator funterm;                 "^
   202            (*get_denominator*)
   203            "      (num::real) = get_numerator funterm;                     "^
   204            (*get_numerator*)
   205            "      (equ::bool) = (denom = (0::real));                       "^
   206            "      (L_L::bool list) = (SubProblem (''PolyEq'',                 "^
   207            "         [''abcFormula'',''degree_2'',''polynomial'',''univariate'',''equation''], "^
   208            "         [''no_met''])                                             "^
   209            "         [BOOL equ, REAL zzz]);                                "^
   210            "      (facs::real) = factors_from_solution L_L;                "^
   211            "      (eql::real) = Take (num_orig / facs);                    "^ 
   212       
   213            "      (eqr::real) = (Try (Rewrite_Set ''ansatz_rls'' False)) eql;  "^
   214       
   215            "      (eq::bool) = Take (eql = eqr);                           "^
   216            (*Maybe possible to use HOLogic.mk_eq ??*)
   217            "      eq = (Try (Rewrite_Set ''equival_trans'' False)) eq;         "^ 
   218       
   219            "      eq = drop_questionmarks eq;                              "^
   220            "      (z1::real) = (rhs (NTH 1 L_L));                          "^
   221            (* 
   222             * prepare equation for a - eq_a
   223             * therefor substitute z with solution 1 - z1
   224             *)
   225            "      (z2::real) = (rhs (NTH 2 L_L));                          "^
   226        
   227            "      (eq_a::bool) = Take eq;                                  "^
   228            "      eq_a = (Substitute [zzz=z1]) eq;                         "^
   229            "      eq_a = (Rewrite_Set ''norm_Rational'' False) eq_a;           "^
   230            "      (sol_a::bool list) =                                     "^
   231            "                 (SubProblem (''Isac'',                           "^
   232            "                              [''univariate'',''equation''],[''no_met''])  "^
   233            "                              [BOOL eq_a, REAL (A::real)]);    "^
   234            "      (a::real) = (rhs(NTH 1 sol_a));                          "^
   235       
   236            "      (eq_b::bool) = Take eq;                                  "^
   237            "      eq_b =  (Substitute [zzz=z2]) eq_b;                      "^
   238            "      eq_b = (Rewrite_Set ''norm_Rational'' False) eq_b;           "^
   239            "      (sol_b::bool list) =                                     "^
   240            "                 (SubProblem (''Isac'',                           "^
   241            "                              [''univariate'',''equation''],[''no_met''])  "^
   242            "                              [BOOL eq_b, REAL (B::real)]);    "^
   243            "      (b::real) = (rhs(NTH 1 sol_b));                          "^
   244       
   245            "      eqr = drop_questionmarks eqr;                            "^
   246            "      (pbz::real) = Take eqr;                                  "^
   247            "      pbz = ((Substitute [A=a, B=b]) pbz);                     "^
   248       
   249            "      pbz = Rewrite ''ruleYZ'' False pbz;                          "^
   250            "      pbz = drop_questionmarks pbz;                            "^
   251       
   252            "      (X_z::bool) = Take (X_z = pbz);                          "^
   253            "      (n_eq::bool) = (Rewrite_Set ''inverse_z'' False) X_z;        "^
   254            "      n_eq = drop_questionmarks n_eq                           "^
   255            "in n_eq)")]
   256 \<close>
   257 (* same error as in         inverse_ztransform2
   258 :partial_function (tailrec) inverse_ztransform3 :: "bool \<Rightarrow> bool"
   259   where
   260 "inverse_ztransform X_eq =                                               
   261 (*([1], Frm), X z = 3 / (z - 1 / 4 + -1 / 8 * (1 / z))*)                            
   262 (let X = Take X_eq;                                                                 
   263 (*([1], Res), ?X' z = 3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z)))*)                    
   264   X' = Rewrite ''ruleZY'' False X;                                                  
   265 (*            ?X' z*)                                                               
   266   (X'_z::real) = lhs X';                                                            
   267 (*            z *)                                                                  
   268   (zzz::real) = argument_in X'_z;                                                   
   269 (*            3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z)))*)                            
   270   (funterm::real) = rhs X';                                                         
   271 (*-----*)                                                                           
   272   (pbz::real) = (SubProblem (''Isac'',                                              
   273     [''partial_fraction'',''rational'',''simplification''],                         
   274     [''simplification'',''of_rationals'',''to_partial_fraction''])                  
   275 (*([2], Res), 4 / (z - 1 / 2) + -4 / (z - -1 / 4)*)                                 
   276     [REAL funterm, REAL zzz]);                                                      
   277 (*-----*)                                                                           
   278 (*([3], Frm), ?X' z = 4 / (z - 1 / 2) + -4 / (z - -1 / 4)*)                         
   279   (pbz_eq::bool) = Take (X'_z = pbz);                                               
   280 (*([3], Res), ?X' z = 4 * (?z / (z - 1 / 2)) + -4 * (?z / (z - -1 / 4))*)           
   281   pbz_eq = Rewrite ''ruleYZ'' False pbz_eq;                                         
   282 (*               4 * (z / (z - 1 / 2)) + -4 * (z / (z - -1 / 4))*)                  
   283   pbz_eq = drop_questionmarks pbz_eq;                                               
   284 (*([4], Frm), X_z = 4 * (z / (z - 1 / 2)) + -4 * (z / (z - -1 / 4))*)               
   285   (X_zeq::bool) = Take (X_z = rhs pbz_eq);                                          
   286 (*([4], Res), X_z = 4 * (1 / 2) ^^^ ?n * ?u [?n] + -4 * (-1 / 4) ^^^ ?n * ?u [?n]*)
   287   n_eq = (Rewrite_Set ''inverse_z'' False) X_zeq;                                   
   288 (*            X_z = 4 * (1 / 2) ^^^ n * u [n] + -4 * (-1 / 4) ^^^ n * u [n]*)
   289   n_eq = drop_questionmarks n_eq                                                    
   290 (*([], Res), X_z = 4 * (1 / 2) ^^^ n * u [n] + -4 * (-1 / 4) ^^^ n * u [n]*)
   291 in n_eq)                                                                            "
   292 *)
   293 setup \<open>KEStore_Elems.add_mets
   294     [Specify.prep_met thy "met_SP_Ztrans_inv_sub" [] Celem.e_metID
   295       (["SignalProcessing", "Z_Transform", "Inverse_sub"], 
   296         [("#Given" ,["filterExpression X_eq"]),
   297           ("#Find"  ,["stepResponse n_eq"])],
   298         {rew_ord'="tless_true", rls'= Rule.e_rls, calc = [],
   299           srls = Rule.Rls {id="srls_partial_fraction", 
   300               preconds = [], rew_ord = ("termlessI",termlessI),
   301               erls = Rule.append_rls "erls_in_srls_partial_fraction" Rule.e_rls
   302                   [(*for asm in NTH_CONS ...*)
   303                     Rule.Calc ("Orderings.ord_class.less",eval_equ "#less_"),
   304                     (*2nd NTH_CONS pushes n+-1 into asms*)
   305                     Rule.Calc("Groups.plus_class.plus", eval_binop "#add_")], 
   306               srls = Rule.Erls, calc = [], errpatts = [],
   307               rules = [Rule.Thm ("NTH_CONS", @{thm NTH_CONS}),
   308                   Rule.Calc ("Groups.plus_class.plus", eval_binop "#add_"),
   309                   Rule.Thm ("NTH_NIL", @{thm NTH_NIL}),
   310                   Rule.Calc ("Tools.lhs", Tools.eval_lhs "eval_lhs_"),
   311                   Rule.Calc ("Tools.rhs", Tools.eval_rhs"eval_rhs_"),
   312                   Rule.Calc ("Atools.argument'_in", eval_argument_in "Atools.argument'_in"),
   313                   Rule.Calc ("Rational.get_denominator", eval_get_denominator "#get_denominator"),
   314                   Rule.Calc ("Rational.get_numerator", eval_get_numerator "#get_numerator"),
   315                   Rule.Calc ("Partial_Fractions.factors_from_solution",
   316                     eval_factors_from_solution "#factors_from_solution"),
   317                   Rule.Calc("Partial_Fractions.drop_questionmarks", eval_drop_questionmarks "#drop_?")],
   318               scr = Rule.EmptyScr},
   319           prls = Rule.e_rls, crls = Rule.e_rls, errpats = [], nrls = norm_Rational},
   320         (*([], Frm), Problem (Isac, [Inverse, Z_Transform, SignalProcessing])*)
   321         "Script InverseZTransform (X_eq::bool) =            "^
   322           (*([1], Frm), X z = 3 / (z - 1 / 4 + -1 / 8 * (1 / z))*)
   323           "(let X = Take X_eq;                                "^
   324           (*([1], Res), ?X' z = 3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z)))*)
   325           "  X' = Rewrite ''ruleZY'' False X;                     "^
   326           (*            ?X' z*)
   327           "  (X'_z::real) = lhs X';                           "^
   328           (*            z *)
   329           "  (zzz::real) = argument_in X'_z;                  "^
   330           (*            3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z)))*)
   331           "  (funterm::real) = rhs X';                        "^
   332 
   333           "  (pbz::real) = (SubProblem (''Isac'',                "^
   334           "    [''partial_fraction'',''rational'',''simplification''],    "^
   335           "    [''simplification'',''of_rationals'',''to_partial_fraction'']) "^
   336           (*([2], Res), 4 / (z - 1 / 2) + -4 / (z - -1 / 4)*)
   337           "    [REAL funterm, REAL zzz]);                     "^
   338 
   339           (*([3], Frm), ?X' z = 4 / (z - 1 / 2) + -4 / (z - -1 / 4)*)
   340           "  (pbz_eq::bool) = Take (X'_z = pbz);              "^
   341           (*([3], Res), ?X' z = 4 * (?z / (z - 1 / 2)) + -4 * (?z / (z - -1 / 4))*)
   342           "  pbz_eq = Rewrite ''ruleYZ'' False pbz_eq;            "^
   343           (*               4 * (z / (z - 1 / 2)) + -4 * (z / (z - -1 / 4))*)
   344           "  pbz_eq = drop_questionmarks pbz_eq;              "^
   345           (*([4], Frm), X_z = 4 * (z / (z - 1 / 2)) + -4 * (z / (z - -1 / 4))*)
   346           "  (X_zeq::bool) = Take (X_z = rhs pbz_eq);         "^
   347           (*([4], Res), X_z = 4 * (1 / 2) ^^^ ?n * ?u [?n] + -4 * (-1 / 4) ^^^ ?n * ?u [?n]*)
   348           "  n_eq = (Rewrite_Set ''inverse_z'' False) X_zeq;      "^
   349           (*            X_z = 4 * (1 / 2) ^^^ n * u [n] + -4 * (-1 / 4) ^^^ n * u [n]*)
   350           "  n_eq = drop_questionmarks n_eq                   "^
   351           (*([], Res), X_z = 4 * (1 / 2) ^^^ n * u [n] + -4 * (-1 / 4) ^^^ n * u [n]*)
   352           "in n_eq)")]
   353 \<close>
   354 
   355 end
   356