src/Tools/isac/Knowledge/Inverse_Z_Transform.thy
author Mathias Lehnfeld <s1210629013@students.fh-hagenberg.at>
Fri, 13 Jun 2014 10:29:06 +0200
changeset 55444 ede4248a827b
parent 55380 7be2ad0e4acb
child 59269 1da53d1540fe
permissions -rwxr-xr-x
last the_generic_context (in fun prep_rls) removed.
neuper@42256
     1
(* Title:  Test_Z_Transform
neuper@42256
     2
   Author: Jan Rocnik
neuper@42256
     3
   (c) copyright due to lincense terms.
neuper@42256
     4
12345678901234567890123456789012345678901234567890123456789012345678901234567890
neuper@42256
     5
        10        20        30        40        50        60        70        80
neuper@42256
     6
*)
neuper@42256
     7
neuper@42290
     8
theory Inverse_Z_Transform imports PolyEq DiffApp Partial_Fractions begin
neuper@42256
     9
neuper@42256
    10
axiomatization where 
neuper@42256
    11
  rule1: "1 = \<delta>[n]" and
neuper@42256
    12
  rule2: "|| z || > 1 ==> z / (z - 1) = u [n]" and
neuper@42256
    13
  rule3: "|| z || < 1 ==> z / (z - 1) = -u [-n - 1]" and 
jan@42367
    14
  rule4: "c * (z / (z - \<alpha>)) = c * \<alpha>^^^n * u [n]" and
neuper@42256
    15
  rule5: "|| z || < || \<alpha> || ==> z / (z - \<alpha>) = -(\<alpha>^^^n) * u [-n - 1]" and
jan@42367
    16
  rule6: "|| z || > 1 ==> z/(z - 1)^^^2 = n * u [n]" (*and
jan@42367
    17
  rule42: "(a * (z/(z-b)) + c * (z/(z-d))) = (a * b^^^n * u [n] + c * d^^^n * u [n])"*)
neuper@42256
    18
neuper@42256
    19
axiomatization where
jan@42367
    20
  ruleZY: "(X z = a / b) = (X' z = a / (z * b))" and
jan@42367
    21
  ruleYZ: "(a/b + c/d) = (a*(z/b) + c*(z/d))" 
neuper@42256
    22
neuper@42279
    23
subsection{*Define the Field Descriptions for the specification*}
neuper@42279
    24
consts
neuper@42279
    25
  filterExpression  :: "bool => una"
neuper@42279
    26
  stepResponse      :: "bool => una"
neuper@42279
    27
jan@42366
    28
jan@42366
    29
ML {*
s1210629013@55444
    30
val inverse_z = prep_rls'(
jan@42366
    31
  Rls {id = "inverse_z", preconds = [], rew_ord = ("dummy_ord",dummy_ord), 
neuper@42451
    32
	  erls = Erls, srls = Erls, calc = [], errpatts = [],
jan@42366
    33
	  rules = 
jan@42367
    34
	   [
jan@42367
    35
    Thm ("rule4",num_str @{thm rule4})
jan@42366
    36
	   ], 
jan@42366
    37
	 scr = EmptyScr}:rls);
jan@42366
    38
*}
jan@42366
    39
jan@42366
    40
jan@42366
    41
text {*store the rule set for math engine*}
jan@42366
    42
neuper@52125
    43
setup {* KEStore_Elems.add_rlss [("inverse_z", (Context.theory_name @{theory}, inverse_z))] *}
jan@42366
    44
neuper@42256
    45
subsection{*Define the Specification*}
neuper@42256
    46
ML {*
neuper@42256
    47
val thy = @{theory};
neuper@42412
    48
*}
s1210629013@55359
    49
setup {* KEStore_Elems.add_pbts
s1210629013@55339
    50
  [(prep_pbt thy "pbl_SP" [] e_pblID (["SignalProcessing"], [], e_rls, NONE, [])),
s1210629013@55339
    51
    (prep_pbt thy "pbl_SP_Ztrans" [] e_pblID
s1210629013@55339
    52
      (["Z_Transform","SignalProcessing"], [], e_rls, NONE, [])),
s1210629013@55339
    53
    (prep_pbt thy "pbl_SP_Ztrans_inv" [] e_pblID
s1210629013@55339
    54
      (["Inverse", "Z_Transform", "SignalProcessing"],
s1210629013@55339
    55
        (*^ capital letter breaks coding standard
s1210629013@55339
    56
          because "inverse" = Const ("Rings.inverse_class.inverse", ..*)
s1210629013@55339
    57
        [("#Given" ,["filterExpression (X_eq::bool)"]),
s1210629013@55339
    58
          ("#Find"  ,["stepResponse (n_eq::bool)"])],
s1210629013@55339
    59
        append_rls "e_rls" e_rls [(*for preds in where_*)], NONE, 
s1210629013@55339
    60
        [["SignalProcessing","Z_Transform","Inverse"]])),
s1210629013@55339
    61
    (prep_pbt thy "pbl_SP_Ztrans_inv" [] e_pblID
s1210629013@55339
    62
      (["Inverse", "Z_Transform", "SignalProcessing"],
s1210629013@55339
    63
        [("#Given" ,["filterExpression X_eq"]),
s1210629013@55339
    64
          ("#Find"  ,["stepResponse n_eq"])],
s1210629013@55339
    65
        append_rls "e_rls" e_rls [(*for preds in where_*)], NONE, 
s1210629013@55339
    66
        [["SignalProcessing","Z_Transform","Inverse"]]))] *}
neuper@42256
    67
neuper@42256
    68
subsection {*Define Name and Signature for the Method*}
neuper@42256
    69
consts
neuper@42256
    70
  InverseZTransform :: "[bool, bool] => bool"
neuper@42256
    71
    ("((Script InverseZTransform (_ =))// (_))" 9)
neuper@42256
    72
neuper@42277
    73
subsection {*Setup Parent Nodes in Hierarchy of Method*}
s1210629013@55380
    74
ML {* val thy = @{theory}; (*latest version of thy required*) *}
s1210629013@55373
    75
setup {* KEStore_Elems.add_mets
s1210629013@55373
    76
  [prep_met thy "met_SP" [] e_metID
s1210629013@55373
    77
      (["SignalProcessing"], [],
s1210629013@55373
    78
        {rew_ord'="tless_true", rls'= e_rls, calc = [], srls = e_rls, prls = e_rls, crls = e_rls,
s1210629013@55373
    79
          errpats = [], nrls = e_rls}, "empty_script"),
s1210629013@55373
    80
    prep_met thy "met_SP_Ztrans" [] e_metID
s1210629013@55373
    81
      (["SignalProcessing", "Z_Transform"], [],
s1210629013@55373
    82
        {rew_ord'="tless_true", rls'= e_rls, calc = [], srls = e_rls, prls = e_rls, crls = e_rls,
s1210629013@55373
    83
          errpats = [], nrls = e_rls}, "empty_script"),
s1210629013@55373
    84
    prep_met thy "met_SP_Ztrans_inv" [] e_metID
s1210629013@55373
    85
      (["SignalProcessing", "Z_Transform", "Inverse"], 
s1210629013@55373
    86
        [("#Given" ,["filterExpression (X_eq::bool)"]),
s1210629013@55373
    87
          ("#Find"  ,["stepResponse (n_eq::bool)"])],
s1210629013@55373
    88
        {rew_ord'="tless_true", rls'= e_rls, calc = [], srls = e_rls, prls = e_rls, crls = e_rls,
s1210629013@55373
    89
          errpats = [], nrls = e_rls},
s1210629013@55373
    90
        "Script InverseZTransform (X_eq::bool) =" ^ (*(1/z) instead of z ^^^ -1*)
s1210629013@55373
    91
          " (let X = Take X_eq;" ^
s1210629013@55373
    92
          "      X' = Rewrite ruleZY False X;" ^ (*z * denominator*)
s1210629013@55373
    93
          "      X' = (Rewrite_Set norm_Rational False) X';" ^ (*simplify*)
s1210629013@55373
    94
          "      funterm = Take (rhs X');" ^ (*drop X' z = for equation solving*)
s1210629013@55373
    95
          "      denom = (Rewrite_Set partial_fraction False) funterm;" ^ (*get_denominator*)
s1210629013@55373
    96
          "      equ = (denom = (0::real));" ^
s1210629013@55373
    97
          "      fun_arg = Take (lhs X');" ^
s1210629013@55373
    98
          "      arg = (Rewrite_Set partial_fraction False) X';" ^ (*get_argument TODO*)
s1210629013@55373
    99
          "      (L_L::bool list) =                                    " ^
s1210629013@55373
   100
          "            (SubProblem (Test',                            " ^
s1210629013@55373
   101
          "                         [LINEAR,univariate,equation,test]," ^
s1210629013@55373
   102
          "                         [Test,solve_linear])              " ^
s1210629013@55373
   103
          "                        [BOOL equ, REAL z])              " ^
s1210629013@55373
   104
          "  in X)"),
s1210629013@55373
   105
    prep_met thy "met_SP_Ztrans_inv" [] e_metID
s1210629013@55373
   106
      (["SignalProcessing", "Z_Transform", "Inverse"], 
s1210629013@55373
   107
        [("#Given" ,["filterExpression X_eq"]),
s1210629013@55373
   108
          ("#Find"  ,["stepResponse n_eq"])],
s1210629013@55373
   109
        {rew_ord'="tless_true", rls'= e_rls, calc = [], srls = srls_partial_fraction, prls = e_rls,
s1210629013@55373
   110
          crls = e_rls, errpats = [], nrls = e_rls},
s1210629013@55373
   111
        "Script InverseZTransform (X_eq::bool) =                        "^
s1210629013@55373
   112
           (*(1/z) instead of z ^^^ -1*)
s1210629013@55373
   113
           "(let X = Take X_eq;                                            "^
s1210629013@55373
   114
           "      X' = Rewrite ruleZY False X;                             "^
s1210629013@55373
   115
           (*z * denominator*)
s1210629013@55373
   116
           "      (num_orig::real) = get_numerator (rhs X');               "^
s1210629013@55373
   117
           "      X' = (Rewrite_Set norm_Rational False) X';               "^
s1210629013@55373
   118
           (*simplify*)
s1210629013@55373
   119
           "      (X'_z::real) = lhs X';                                   "^
s1210629013@55373
   120
           "      (zzz::real) = argument_in X'_z;                          "^
s1210629013@55373
   121
           "      (funterm::real) = rhs X';                                "^
s1210629013@55373
   122
           (*drop X' z = for equation solving*)
s1210629013@55373
   123
           "      (denom::real) = get_denominator funterm;                 "^
s1210629013@55373
   124
           (*get_denominator*)
s1210629013@55373
   125
           "      (num::real) = get_numerator funterm;                     "^
s1210629013@55373
   126
           (*get_numerator*)
s1210629013@55373
   127
           "      (equ::bool) = (denom = (0::real));                       "^
s1210629013@55373
   128
           "      (L_L::bool list) = (SubProblem (PolyEq',                 "^
s1210629013@55373
   129
           "         [abcFormula,degree_2,polynomial,univariate,equation], "^
s1210629013@55373
   130
           "         [no_met])                                             "^
s1210629013@55373
   131
           "         [BOOL equ, REAL zzz]);                                "^
s1210629013@55373
   132
           "      (facs::real) = factors_from_solution L_L;                "^
s1210629013@55373
   133
           "      (eql::real) = Take (num_orig / facs);                    "^ 
s1210629013@55373
   134
      
s1210629013@55373
   135
           "      (eqr::real) = (Try (Rewrite_Set ansatz_rls False)) eql;  "^
s1210629013@55373
   136
      
s1210629013@55373
   137
           "      (eq::bool) = Take (eql = eqr);                           "^
s1210629013@55373
   138
           (*Maybe possible to use HOLogic.mk_eq ??*)
s1210629013@55373
   139
           "      eq = (Try (Rewrite_Set equival_trans False)) eq;         "^ 
s1210629013@55373
   140
      
s1210629013@55373
   141
           "      eq = drop_questionmarks eq;                              "^
s1210629013@55373
   142
           "      (z1::real) = (rhs (NTH 1 L_L));                          "^
s1210629013@55373
   143
           (* 
s1210629013@55373
   144
            * prepare equation for a - eq_a
s1210629013@55373
   145
            * therefor substitute z with solution 1 - z1
s1210629013@55373
   146
            *)
s1210629013@55373
   147
           "      (z2::real) = (rhs (NTH 2 L_L));                          "^
s1210629013@55373
   148
       
s1210629013@55373
   149
           "      (eq_a::bool) = Take eq;                                  "^
s1210629013@55373
   150
           "      eq_a = (Substitute [zzz=z1]) eq;                         "^
s1210629013@55373
   151
           "      eq_a = (Rewrite_Set norm_Rational False) eq_a;           "^
s1210629013@55373
   152
           "      (sol_a::bool list) =                                     "^
s1210629013@55373
   153
           "                 (SubProblem (Isac',                           "^
s1210629013@55373
   154
           "                              [univariate,equation],[no_met])  "^
s1210629013@55373
   155
           "                              [BOOL eq_a, REAL (A::real)]);    "^
s1210629013@55373
   156
           "      (a::real) = (rhs(NTH 1 sol_a));                          "^
s1210629013@55373
   157
      
s1210629013@55373
   158
           "      (eq_b::bool) = Take eq;                                  "^
s1210629013@55373
   159
           "      eq_b =  (Substitute [zzz=z2]) eq_b;                      "^
s1210629013@55373
   160
           "      eq_b = (Rewrite_Set norm_Rational False) eq_b;           "^
s1210629013@55373
   161
           "      (sol_b::bool list) =                                     "^
s1210629013@55373
   162
           "                 (SubProblem (Isac',                           "^
s1210629013@55373
   163
           "                              [univariate,equation],[no_met])  "^
s1210629013@55373
   164
           "                              [BOOL eq_b, REAL (B::real)]);    "^
s1210629013@55373
   165
           "      (b::real) = (rhs(NTH 1 sol_b));                          "^
s1210629013@55373
   166
      
s1210629013@55373
   167
           "      eqr = drop_questionmarks eqr;                            "^
s1210629013@55373
   168
           "      (pbz::real) = Take eqr;                                  "^
s1210629013@55373
   169
           "      pbz = ((Substitute [A=a, B=b]) pbz);                     "^
s1210629013@55373
   170
      
s1210629013@55373
   171
           "      pbz = Rewrite ruleYZ False pbz;                          "^
s1210629013@55373
   172
           "      pbz = drop_questionmarks pbz;                            "^
s1210629013@55373
   173
      
s1210629013@55373
   174
           "      (X_z::bool) = Take (X_z = pbz);                          "^
s1210629013@55373
   175
           "      (n_eq::bool) = (Rewrite_Set inverse_z False) X_z;        "^
s1210629013@55373
   176
           "      n_eq = drop_questionmarks n_eq                           "^
s1210629013@55373
   177
           "in n_eq)"),
s1210629013@55373
   178
    prep_met thy "met_SP_Ztrans_inv_sub" [] e_metID
s1210629013@55373
   179
      (["SignalProcessing", "Z_Transform", "Inverse_sub"], 
s1210629013@55373
   180
        [("#Given" ,["filterExpression X_eq"]),
s1210629013@55373
   181
          ("#Find"  ,["stepResponse n_eq"])],
s1210629013@55373
   182
        {rew_ord'="tless_true", rls'= e_rls, calc = [],
s1210629013@55373
   183
          srls = Rls {id="srls_partial_fraction", 
s1210629013@55373
   184
              preconds = [], rew_ord = ("termlessI",termlessI),
s1210629013@55373
   185
              erls = append_rls "erls_in_srls_partial_fraction" e_rls
s1210629013@55373
   186
                  [(*for asm in NTH_CONS ...*)
s1210629013@55373
   187
                    Calc ("Orderings.ord_class.less",eval_equ "#less_"),
s1210629013@55373
   188
                    (*2nd NTH_CONS pushes n+-1 into asms*)
s1210629013@55373
   189
                    Calc("Groups.plus_class.plus", eval_binop "#add_")], 
s1210629013@55373
   190
              srls = Erls, calc = [], errpatts = [],
s1210629013@55373
   191
              rules = [Thm ("NTH_CONS",num_str @{thm NTH_CONS}),
s1210629013@55373
   192
                  Calc ("Groups.plus_class.plus", eval_binop "#add_"),
s1210629013@55373
   193
                  Thm ("NTH_NIL",num_str @{thm NTH_NIL}),
s1210629013@55373
   194
                  Calc ("Tools.lhs", eval_lhs "eval_lhs_"),
s1210629013@55373
   195
                  Calc ("Tools.rhs", eval_rhs"eval_rhs_"),
s1210629013@55373
   196
                  Calc ("Atools.argument'_in", eval_argument_in "Atools.argument'_in"),
s1210629013@55373
   197
                  Calc ("Rational.get_denominator", eval_get_denominator "#get_denominator"),
s1210629013@55373
   198
                  Calc ("Rational.get_numerator", eval_get_numerator "#get_numerator"),
s1210629013@55373
   199
                  Calc ("Partial_Fractions.factors_from_solution",
s1210629013@55373
   200
                    eval_factors_from_solution "#factors_from_solution"),
s1210629013@55373
   201
                  Calc("Partial_Fractions.drop_questionmarks", eval_drop_questionmarks "#drop_?")],
s1210629013@55373
   202
              scr = EmptyScr},
s1210629013@55373
   203
          prls = e_rls, crls = e_rls, errpats = [], nrls = norm_Rational},
s1210629013@55373
   204
        (*([], Frm), Problem (Isac, [Inverse, Z_Transform, SignalProcessing])*)
s1210629013@55373
   205
        "Script InverseZTransform (X_eq::bool) =            "^
s1210629013@55373
   206
          (*([1], Frm), X z = 3 / (z - 1 / 4 + -1 / 8 * (1 / z))*)
s1210629013@55373
   207
          "(let X = Take X_eq;                                "^
s1210629013@55373
   208
          (*([1], Res), ?X' z = 3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z)))*)
s1210629013@55373
   209
          "  X' = Rewrite ruleZY False X;                     "^
s1210629013@55373
   210
          (*            ?X' z*)
s1210629013@55373
   211
          "  (X'_z::real) = lhs X';                           "^
s1210629013@55373
   212
          (*            z *)
s1210629013@55373
   213
          "  (zzz::real) = argument_in X'_z;                  "^
s1210629013@55373
   214
          (*            3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z)))*)
s1210629013@55373
   215
          "  (funterm::real) = rhs X';                        "^
s1210629013@55373
   216
s1210629013@55373
   217
          "  (pbz::real) = (SubProblem (Isac',                "^
s1210629013@55373
   218
          "    [partial_fraction,rational,simplification],    "^
s1210629013@55373
   219
          "    [simplification,of_rationals,to_partial_fraction]) "^
s1210629013@55373
   220
          (*([2], Res), 4 / (z - 1 / 2) + -4 / (z - -1 / 4)*)
s1210629013@55373
   221
          "    [REAL funterm, REAL zzz]);                     "^
s1210629013@55373
   222
s1210629013@55373
   223
          (*([3], Frm), ?X' z = 4 / (z - 1 / 2) + -4 / (z - -1 / 4)*)
s1210629013@55373
   224
          "  (pbz_eq::bool) = Take (X'_z = pbz);              "^
s1210629013@55373
   225
          (*([3], Res), ?X' z = 4 * (?z / (z - 1 / 2)) + -4 * (?z / (z - -1 / 4))*)
s1210629013@55373
   226
          "  pbz_eq = Rewrite ruleYZ False pbz_eq;            "^
s1210629013@55373
   227
          (*               4 * (z / (z - 1 / 2)) + -4 * (z / (z - -1 / 4))*)
s1210629013@55373
   228
          "  pbz_eq = drop_questionmarks pbz_eq;              "^
s1210629013@55373
   229
          (*([4], Frm), X_z = 4 * (z / (z - 1 / 2)) + -4 * (z / (z - -1 / 4))*)
s1210629013@55373
   230
          "  (X_zeq::bool) = Take (X_z = rhs pbz_eq);         "^
s1210629013@55373
   231
          (*([4], Res), X_z = 4 * (1 / 2) ^^^ ?n * ?u [?n] + -4 * (-1 / 4) ^^^ ?n * ?u [?n]*)
s1210629013@55373
   232
          "  n_eq = (Rewrite_Set inverse_z False) X_zeq;      "^
s1210629013@55373
   233
          (*            X_z = 4 * (1 / 2) ^^^ n * u [n] + -4 * (-1 / 4) ^^^ n * u [n]*)
s1210629013@55373
   234
          "  n_eq = drop_questionmarks n_eq                   "^
s1210629013@55373
   235
          (*([], Res), X_z = 4 * (1 / 2) ^^^ n * u [n] + -4 * (-1 / 4) ^^^ n * u [n]*)
s1210629013@55373
   236
          "in n_eq)")]
s1210629013@55373
   237
*}
neuper@42256
   238
neuper@42256
   239
end
neuper@42256
   240