src/Tools/isac/Knowledge/Inverse_Z_Transform.thy
author Walther Neuper <neuper@ist.tugraz.at>
Sun, 22 Sep 2013 18:09:05 +0200
changeset 52125 6f1d3415dc68
parent 42451 bc03b5d60547
child 52155 e4ddf21390fd
permissions -rwxr-xr-x
add functions accessing Theory_Data in parallel to those accessing "ruleset' = Unsynchronized.ref"

updates have been done incrementally following Build_Isac.thy:
# ./bin/isabelle jedit -l HOL src/Tools/isac/ProgLang/ProgLang.thy &
# ./bin/isabelle jedit -l HOL src/Tools/isac/Interpret/Interpret.thy &
# ./bin/isabelle jedit -l HOL src/Tools/isac/xmlsrc/xmlsrc.thy &
# ./bin/isabelle jedit -l HOL src/Tools/isac/Frontend/Frontend.thy &

Note, that the original access function "fun assoc_rls" is still outcommented;
so the old and new functionality is established in parallel.
     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{*Define the Field Descriptions for the specification*}
    24 consts
    25   filterExpression  :: "bool => una"
    26   stepResponse      :: "bool => una"
    27 
    28 
    29 ML {*
    30 val inverse_z = prep_rls(
    31   Rls {id = "inverse_z", preconds = [], rew_ord = ("dummy_ord",dummy_ord), 
    32 	  erls = Erls, srls = Erls, calc = [], errpatts = [],
    33 	  rules = 
    34 	   [
    35     Thm ("rule4",num_str @{thm rule4})
    36 	   ], 
    37 	 scr = EmptyScr}:rls);
    38 *}
    39 
    40 
    41 text {*store the rule set for math engine*}
    42 
    43 ML {*
    44 ruleset' := overwritelthy @{theory} (!ruleset',
    45   [("inverse_z", inverse_z)
    46    ]);
    47 *}
    48 setup {* KEStore_Elems.add_rlss [("inverse_z", (Context.theory_name @{theory}, inverse_z))] *}
    49 
    50 subsection{*Define the Specification*}
    51 ML {*
    52 val thy = @{theory};
    53 
    54 store_pbt
    55  (prep_pbt thy "pbl_SP" [] e_pblID
    56  (["SignalProcessing"], [], e_rls, NONE, []));
    57 store_pbt
    58  (prep_pbt thy "pbl_SP_Ztrans" [] e_pblID
    59  (["Z_Transform","SignalProcessing"], [], e_rls, NONE, []));
    60 store_pbt
    61  (prep_pbt thy "pbl_SP_Ztrans_inv" [] e_pblID
    62  (["Inverse", "Z_Transform", "SignalProcessing"],
    63   (*^ capital letter breaks coding standard
    64     because "inverse" = Const ("Rings.inverse_class.inverse", ..*)
    65   [("#Given" ,["filterExpression (X_eq::bool)"]),
    66    ("#Find"  ,["stepResponse (n_eq::bool)"])
    67   ],
    68   append_rls "e_rls" e_rls [(*for preds in where_*)], NONE, 
    69   [["SignalProcessing","Z_Transform","Inverse"]]));
    70 *}
    71 ML {*
    72   store_pbt
    73    (prep_pbt thy "pbl_SP_Ztrans_inv" [] e_pblID
    74    (["Inverse", "Z_Transform", "SignalProcessing"],
    75     [("#Given" ,["filterExpression X_eq"]),
    76      ("#Find"  ,["stepResponse n_eq"])
    77     ],
    78     append_rls "e_rls" e_rls [(*for preds in where_*)], NONE, 
    79     [["SignalProcessing","Z_Transform","Inverse"]]));
    80 *}
    81 
    82 subsection {*Define Name and Signature for the Method*}
    83 consts
    84   InverseZTransform :: "[bool, bool] => bool"
    85     ("((Script InverseZTransform (_ =))// (_))" 9)
    86 
    87 subsection {*Setup Parent Nodes in Hierarchy of Method*}
    88 ML {*
    89 store_met
    90  (prep_met thy "met_SP" [] e_metID
    91  (["SignalProcessing"], [],
    92    {rew_ord'="tless_true", rls'= e_rls, calc = [], srls = e_rls, prls = e_rls,
    93     crls = e_rls, errpats = [], nrls = e_rls}, "empty_script"));
    94 store_met
    95  (prep_met thy "met_SP_Ztrans" [] e_metID
    96  (["SignalProcessing", "Z_Transform"], [],
    97    {rew_ord'="tless_true", rls'= e_rls, calc = [], srls = e_rls, prls = e_rls,
    98     crls = e_rls, errpats = [], nrls = e_rls}, "empty_script"));
    99 val thy = @{theory}; (*latest version of thy required*)
   100 store_met
   101  (prep_met thy "met_SP_Ztrans_inv" [] e_metID
   102  (["SignalProcessing", "Z_Transform", "Inverse"], 
   103   [("#Given" ,["filterExpression (X_eq::bool)"]),
   104    ("#Find"  ,["stepResponse (n_eq::bool)"])
   105   ],
   106    {rew_ord'="tless_true", rls'= e_rls, calc = [], srls = e_rls, prls = e_rls,
   107     crls = e_rls, errpats = [], nrls = 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  ));
   124 *}
   125 ML {*
   126   store_met(
   127     prep_met thy "met_SP_Ztrans_inv" [] e_metID
   128     (["SignalProcessing",
   129       "Z_Transform",
   130       "Inverse"], 
   131      [
   132        ("#Given" ,["filterExpression X_eq"]),
   133        ("#Find"  ,["stepResponse n_eq"])
   134      ],
   135      {
   136        rew_ord'="tless_true",
   137        rls'= e_rls, calc = [],
   138        srls = srls_partial_fraction, 
   139        prls = e_rls,
   140        crls = e_rls, errpats = [], nrls = e_rls
   141      },
   142      "Script InverseZTransform (X_eq::bool) =                        "^
   143      (*(1/z) instead of z ^^^ -1*)
   144      "(let X = Take X_eq;                                            "^
   145      "      X' = Rewrite ruleZY False X;                             "^
   146      (*z * denominator*)
   147      "      (num_orig::real) = get_numerator (rhs X');               "^
   148      "      X' = (Rewrite_Set norm_Rational False) X';               "^
   149      (*simplify*)
   150      "      (X'_z::real) = lhs X';                                   "^
   151      "      (zzz::real) = argument_in X'_z;                          "^
   152      "      (funterm::real) = rhs X';                                "^
   153      (*drop X' z = for equation solving*)
   154      "      (denom::real) = get_denominator funterm;                 "^
   155      (*get_denominator*)
   156      "      (num::real) = get_numerator funterm;                     "^
   157      (*get_numerator*)
   158      "      (equ::bool) = (denom = (0::real));                       "^
   159      "      (L_L::bool list) = (SubProblem (PolyEq',                 "^
   160      "         [abcFormula,degree_2,polynomial,univariate,equation], "^
   161      "         [no_met])                                             "^
   162      "         [BOOL equ, REAL zzz]);                                "^
   163      "      (facs::real) = factors_from_solution L_L;                "^
   164      "      (eql::real) = Take (num_orig / facs);                    "^ 
   165 
   166      "      (eqr::real) = (Try (Rewrite_Set ansatz_rls False)) eql;  "^
   167 
   168      "      (eq::bool) = Take (eql = eqr);                           "^
   169      (*Maybe possible to use HOLogic.mk_eq ??*)
   170      "      eq = (Try (Rewrite_Set equival_trans False)) eq;         "^ 
   171 
   172      "      eq = drop_questionmarks eq;                              "^
   173      "      (z1::real) = (rhs (NTH 1 L_L));                          "^
   174      (* 
   175       * prepare equation for a - eq_a
   176       * therefor substitute z with solution 1 - z1
   177       *)
   178      "      (z2::real) = (rhs (NTH 2 L_L));                          "^
   179  
   180      "      (eq_a::bool) = Take eq;                                  "^
   181      "      eq_a = (Substitute [zzz=z1]) eq;                         "^
   182      "      eq_a = (Rewrite_Set norm_Rational False) eq_a;           "^
   183      "      (sol_a::bool list) =                                     "^
   184      "                 (SubProblem (Isac',                           "^
   185      "                              [univariate,equation],[no_met])  "^
   186      "                              [BOOL eq_a, REAL (A::real)]);    "^
   187      "      (a::real) = (rhs(NTH 1 sol_a));                          "^
   188 
   189      "      (eq_b::bool) = Take eq;                                  "^
   190      "      eq_b =  (Substitute [zzz=z2]) eq_b;                      "^
   191      "      eq_b = (Rewrite_Set norm_Rational False) eq_b;           "^
   192      "      (sol_b::bool list) =                                     "^
   193      "                 (SubProblem (Isac',                           "^
   194      "                              [univariate,equation],[no_met])  "^
   195      "                              [BOOL eq_b, REAL (B::real)]);    "^
   196      "      (b::real) = (rhs(NTH 1 sol_b));                          "^
   197 
   198      "      eqr = drop_questionmarks eqr;                            "^
   199      "      (pbz::real) = Take eqr;                                  "^
   200      "      pbz = ((Substitute [A=a, B=b]) pbz);                     "^
   201 
   202      "      pbz = Rewrite ruleYZ False pbz;                          "^
   203      "      pbz = drop_questionmarks pbz;                            "^
   204 
   205      "      (X_z::bool) = Take (X_z = pbz);                          "^
   206      "      (n_eq::bool) = (Rewrite_Set inverse_z False) X_z;        "^
   207      "      n_eq = drop_questionmarks n_eq                           "^
   208      "in n_eq)" 
   209     )
   210            );
   211 
   212 store_met (prep_met thy "met_SP_Ztrans_inv_sub" [] e_metID
   213   (["SignalProcessing", "Z_Transform", "Inverse_sub"], 
   214    [("#Given" ,["filterExpression X_eq"]),
   215     ("#Find"  ,["stepResponse n_eq"])],
   216    {rew_ord'="tless_true",
   217     rls'= e_rls, calc = [],
   218     srls = Rls {id="srls_partial_fraction", 
   219       preconds = [],
   220       rew_ord = ("termlessI",termlessI),
   221       erls = append_rls "erls_in_srls_partial_fraction" e_rls
   222        [(*for asm in NTH_CONS ...*)
   223         Calc ("Orderings.ord_class.less",eval_equ "#less_"),
   224         (*2nd NTH_CONS pushes n+-1 into asms*)
   225         Calc("Groups.plus_class.plus", eval_binop "#add_")], 
   226         srls = Erls, calc = [], errpatts = [],
   227         rules = [
   228           Thm ("NTH_CONS",num_str @{thm NTH_CONS}),
   229           Calc("Groups.plus_class.plus", eval_binop "#add_"),
   230           Thm ("NTH_NIL",num_str @{thm NTH_NIL}),
   231           Calc("Tools.lhs", eval_lhs "eval_lhs_"),
   232           Calc("Tools.rhs", eval_rhs"eval_rhs_"),
   233           Calc("Atools.argument'_in", eval_argument_in "Atools.argument'_in"),
   234           Calc("Rational.get_denominator", eval_get_denominator "#get_denominator"),
   235           Calc("Rational.get_numerator", eval_get_numerator "#get_numerator"),
   236           Calc("Partial_Fractions.factors_from_solution",
   237             eval_factors_from_solution "#factors_from_solution"),
   238           Calc("Partial_Fractions.drop_questionmarks", eval_drop_questionmarks "#drop_?")],
   239           scr = EmptyScr},
   240     prls = e_rls, crls = e_rls, errpats = [], nrls = norm_Rational},
   241    "Script InverseZTransform (X_eq::bool) =            "^(*([], Frm), Problem (Isac, [Inverse, Z_Transform, SignalProcessing])*)
   242    "(let X = Take X_eq;                                "^(*([1], Frm), X z = 3 / (z - 1 / 4 + -1 / 8 * (1 / z))*)
   243    "  X' = Rewrite ruleZY False X;                     "^(*([1], Res), ?X' z = 3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z)))*)
   244    "  (X'_z::real) = lhs X';                           "^(*            ?X' z*)
   245    "  (zzz::real) = argument_in X'_z;                  "^(*            z *)
   246    "  (funterm::real) = rhs X';                        "^(*            3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z)))*)
   247 
   248    "  (pbz::real) = (SubProblem (Isac',                "^(**)
   249    "    [partial_fraction,rational,simplification],    "^
   250    "    [simplification,of_rationals,to_partial_fraction]) "^
   251    "    [REAL funterm, REAL zzz]);                     "^(*([2], Res), 4 / (z - 1 / 2) + -4 / (z - -1 / 4)*)
   252 
   253    "  (pbz_eq::bool) = Take (X'_z = pbz);              "^(*([3], Frm), ?X' z = 4 / (z - 1 / 2) + -4 / (z - -1 / 4)*)
   254    "  pbz_eq = Rewrite ruleYZ False pbz_eq;            "^(*([3], Res), ?X' z = 4 * (?z / (z - 1 / 2)) + -4 * (?z / (z - -1 / 4))*)
   255    "  pbz_eq = drop_questionmarks pbz_eq;              "^(*               4 * (z / (z - 1 / 2)) + -4 * (z / (z - -1 / 4))*)
   256    "  (X_zeq::bool) = Take (X_z = rhs pbz_eq);         "^(*([4], Frm), X_z = 4 * (z / (z - 1 / 2)) + -4 * (z / (z - -1 / 4))*)
   257    "  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]*)
   258    "  n_eq = drop_questionmarks n_eq                   "^(*            X_z = 4 * (1 / 2) ^^^ n * u [n] + -4 * (-1 / 4) ^^^ n * u [n]*)
   259    "in n_eq)"                                            (*([], Res), X_z = 4 * (1 / 2) ^^^ n * u [n] + -4 * (-1 / 4) ^^^ n * u [n]*)
   260   ));
   261 
   262 *}
   263 
   264 end
   265