test/Tools/isac/ADDTESTS/course/SignalProcess/Build_Inverse_Z_Transform.thy
author Walther Neuper <neuper@ist.tugraz.at>
Thu, 13 Oct 2011 15:03:28 +0200
branchdecompose-isar
changeset 42315 c2e6ac4a5d04
parent 42310 55931ca19f4d
child 42335 7bb5070f2415
permissions -rwxr-xr-x
Build_Inverse_Z_Transform 1 step further

the expected error in isac:
prep. repair Apply_Method without init_form
was a simple error in the CTP-based program
     1 (* Title:  Build_Inverse_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 Build_Inverse_Z_Transform imports Isac
     9   
    10 begin
    11 
    12 text{* We stepwise build Inverse_Z_Transform.thy as an exercise.
    13   Because subsection "Stepwise Check the Program" requires 
    14   Inverse_Z_Transform.thy as a subtheory of Isac.thy, the setup has been changed 
    15   from "theory Inverse_Z_Transform imports Isac begin.." to the above.
    16 
    17   ATTENTION WITH NAMES OF IDENTIFIERS WHEN GOING INTO INTERNALS:
    18   Here in this theory there are the internal names twice, for instance we have
    19   (Thm.derivation_name @{thm rule1} = "Build_Inverse_Z_Transform.rule1") = true;
    20   but actually in us will be "Inverse_Z_Transform.rule1"
    21 *}
    22 ML {*val thy = @{theory Isac};*}
    23 
    24 
    25 section {*trials towards Z transform *}
    26 text{*===============================*}
    27 subsection {*terms*}
    28 ML {*
    29 @{term "1 < || z ||"};
    30 @{term "z / (z - 1)"};
    31 @{term "-u -n - 1"};
    32 @{term "-u [-n - 1]"}; (*[ ] denotes lists !!!*)
    33 @{term "z /(z - 1) = -u [-n - 1]"};Isac
    34 @{term "1 < || z || ==> z / (z - 1) = -u [-n - 1]"};
    35 term2str @{term "1 < || z || ==> z / (z - 1) = -u [-n - 1]"};
    36 *}
    37 ML {*
    38 (*alpha -->  "</alpha>" *)
    39 @{term "\<alpha> "};
    40 @{term "\<delta> "};
    41 @{term "\<phi> "};
    42 @{term "\<rho> "};
    43 term2str @{term "\<rho> "};
    44 *}
    45 
    46 subsection {*rules*}
    47 (*axiomatization "z / (z - 1) = -u [-n - 1]" Illegal variable name: "z / (z - 1) = -u [-n - 1]" *)
    48 (*definition     "z / (z - 1) = -u [-n - 1]" Bad head of lhs: existing constant "op /"*)
    49 axiomatization where 
    50   rule1: "1 = \<delta>[n]" and
    51   rule2: "|| z || > 1 ==> z / (z - 1) = u [n]" and
    52   rule3: "|| z || < 1 ==> z / (z - 1) = -u [-n - 1]" and 
    53   rule4: "|| z || > || \<alpha> || ==> z / (z - \<alpha>) = \<alpha>^^^n * u [n]" and
    54   rule5: "|| z || < || \<alpha> || ==> z / (z - \<alpha>) = -(\<alpha>^^^n) * u [-n - 1]" and
    55   rule6: "|| z || > 1 ==> z/(z - 1)^^^2 = n * u [n]"
    56 ML {*
    57 @{thm rule1};
    58 @{thm rule2};
    59 @{thm rule3};
    60 @{thm rule4};
    61 *}
    62 
    63 subsection {*apply rules*}
    64 ML {*
    65 val inverse_Z = append_rls "inverse_Z" e_rls
    66   [ Thm  ("rule3",num_str @{thm rule3}),
    67     Thm  ("rule4",num_str @{thm rule4}),
    68     Thm  ("rule1",num_str @{thm rule1})   
    69   ];
    70 
    71 val t = str2term "z / (z - 1) + z / (z - \<alpha>) + 1";
    72 val SOME (t', asm) = rewrite_set_ thy true inverse_Z t;
    73 term2str t' = "z / (z - ?\<delta> [?n]) + z / (z - \<alpha>) + ?\<delta> [?n]"; (*attention rule1 !!!*)
    74 *}
    75 ML {*
    76 val (thy, ro, er) = (@{theory Isac}, tless_true, eval_rls);
    77 *}
    78 ML {*
    79 val SOME (t, asm1) = rewrite_ thy ro er true (num_str @{thm rule3}) t;
    80 term2str t = "- ?u [- ?n - 1] + z / (z - \<alpha>) + 1"; (*- real *)
    81 term2str t;*}
    82 ML {*
    83 val SOME (t, asm2) = rewrite_ thy ro er true (num_str @{thm rule4}) t;
    84 term2str t = "- ?u [- ?n - 1] + \<alpha> ^^^ ?n * ?u [?n] + 1"; (*- real *)
    85 term2str t;
    86 *}
    87 ML {*
    88 val SOME (t, asm3) = rewrite_ thy ro er true (num_str @{thm rule1}) t;
    89 term2str t = "- ?u [- ?n - 1] + \<alpha> ^^^ ?n * ?u [?n] + ?\<delta> [?n]"; (*- real *)
    90 term2str t;
    91 *}
    92 ML {*
    93 terms2str (asm1 @ asm2 @ asm3);
    94 *}
    95 
    96 section {*Prepare steps for CTP-based programming language*}
    97 text{*TODO insert Calculation (Referenz?!)
    98 
    99 The goal... realized in sections below, in Sect.\ref{spec-meth} and Sect.\ref{prog-steps} 
   100 
   101 the reader is advised to jump between the subsequent subsections and the respective steps in Sect.\ref{prog-steps} 
   102 
   103 *}
   104 subsection {*prepare expression \label{prep-expr}*}
   105 ML {*
   106 val ctxt = ProofContext.init_global @{theory Isac};
   107 val ctxt = declare_constraints' [@{term "z::real"}] ctxt;
   108 
   109 val SOME fun1 = parseNEW ctxt "X z = 3 / (z - 1/4 + -1/8 * z ^^^ -1)"; term2str fun1;
   110 val SOME fun1' = parseNEW ctxt "X z = 3 / (z - 1/4 + -1/8 * (1/z))"; term2str fun1';
   111 *}
   112 
   113 subsubsection {*multply with z*}
   114 axiomatization where
   115   ruleZY: "(X z = a / b) = (X' z = a / (z * b))"
   116 
   117 ML {*
   118 val (thy, ro, er) = (@{theory Isac}, tless_true, eval_rls);
   119 val SOME (fun2, asm1) = rewrite_ thy ro er true  @{thm ruleZY} fun1; term2str fun2;
   120 val SOME (fun2', asm1) = rewrite_ thy ro er true  @{thm ruleZY} fun1'; term2str fun2';
   121 
   122 val SOME (fun3,_) = rewrite_set_ @{theory Isac} false norm_Rational fun2;
   123 term2str fun3; (*fails on x^^^(-1) TODO*)
   124 val SOME (fun3',_) = rewrite_set_ @{theory Isac} false norm_Rational fun2';
   125 term2str fun3'; (*OK*)
   126 *}
   127 
   128 subsubsection {*get argument of X': z is the variable the equation is solved for*}
   129 text{*grep... Atools.thy, Tools.thy contain general utilities: eval_argument_in, eval_rhs, eval_lhs,...
   130 
   131 grep -r "fun eva_" ... shows all functions witch can be used in a script.
   132 lookup this files how to build and handle such functions.
   133 
   134 the next section shows how to introduce such a function.
   135 *}
   136 
   137 subsubsection {*Decompose given term into lhs = rhs*}
   138 ML {*
   139   val (_, expr) = HOLogic.dest_eq fun3'; term2str expr;
   140   val (_, denom) = HOLogic.dest_bin "Rings.inverse_class.divide" (type_of expr) expr;
   141   term2str denom = "-1 + -2 * z + 8 * z ^^^ 2";
   142 *}
   143 text {*we have rhs in the Script language, but we need a function 
   144   which gets the denominator of a fraction*}
   145 
   146 text{*---------------------------begin partial fractions snip--------------------------*}
   147 
   148 subsubsection {*get the denominator out of a fraction*}
   149 text {*get denominator should become a constant for the isabelle parser: *}
   150 
   151 consts
   152   get_denominator :: "real => real"
   153 
   154 text {* With the above definition we run into problems with parsing the Script InverseZTransform:
   155   This leads to "ambiguous parse trees" and we avoid this by shifting the definition
   156   to Rationa.thy and re-building Isac.
   157   ATTENTION: from now on Build_Inverse_Z_Transform mimics a build from scratch;
   158   it only works due to re-building Isac several times (indicated explicityl).
   159 *}
   160 
   161 ML {*
   162 (*("get_denominator", ("Rational.get_denominator", eval_get_denominator ""))*)
   163 fun eval_get_denominator (thmid:string) _ 
   164 		      (t as Const ("Rational.get_denominator", _) $
   165               (Const ("Rings.inverse_class.divide", _) $ num $
   166                 denom)) thy = 
   167         SOME (mk_thmid thmid "" 
   168             (Print_Mode.setmp [] (Syntax.string_of_term (thy2ctxt thy)) denom) "", 
   169 	          Trueprop $ (mk_equality (t, denom)))
   170   | eval_get_denominator _ _ _ _ = NONE; 
   171 
   172 *}
   173 text {* tests of eval_get_denominator see test/Knowledge/rational.sml*}
   174 
   175 
   176 text{*---------------------------end partial fractions snip--------------------------*}
   177 
   178 subsection {*solve equation*}
   179 text {*this type of equation if too general for the present program*}
   180 ML {*
   181 "----------- Minisubplb/100-init-rootp (*OK*)bl.sml ---------------------";
   182 val denominator = parseNEW ctxt "z^^^2 - 1/4*z - 1/8 = 0";
   183 val fmz = ["equality (z^^^2 - 1/4*z - 1/8 = (0::real))", "solveFor z","solutions L"];
   184 val (dI',pI',mI') =("Isac", ["univariate","equation"], ["no_met"]);
   185 (*                           ^^^^^^^^^^^^^^^^^^^^^^ TODO: ISAC determines type of eq*)
   186 *}
   187 text {*Does the Equation Match the Specification ?*}
   188 ML {*
   189 match_pbl fmz (get_pbt ["univariate","equation"]);
   190 *}
   191 ML {*Context.theory_name thy = "Isac"(*==================================================*)*}
   192 
   193 ML {*
   194 val denominator = parseNEW ctxt "-1 + -2 * z + 8 * z ^^^ 2 = 0";
   195 val fmz =                                            (*specification*)
   196   ["equality (-1 + -2 * z + 8 * z ^^^ 2 = (0::real))", (*equality*)
   197    "solveFor z",                                     (*bound variable*)
   198    "solutions L"];                                   (*identifier for solution*)
   199 
   200 val (dI',pI',mI') =
   201   ("Isac", ["abcFormula","degree_2","polynomial","univariate","equation"], ["no_met"]);
   202 *}
   203 text {*Does the Other Equation Match the Specification ?*}
   204 ML {*
   205 match_pbl fmz (get_pbt ["abcFormula","degree_2","polynomial","univariate","equation"]);
   206 *}
   207 text {*Solve Equation Stepwise*}
   208 ML {*
   209 *}
   210 ML {*
   211 val (p,_,f,nxt,_,pt) = CalcTreeTEST [(fmz, (dI',pI',mI'))];
   212 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   213 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   214 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   215 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   216 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   217 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   218 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   219 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   220 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   221 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   222 val (p,_,f,nxt,_,pt) = me nxt p [] pt;         
   223 val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt =..,Check_elementwise "Assumptions")*)
   224 val (p,_,f,nxt,_,pt) = me nxt p [] pt;         
   225 val (p,_,f,nxt,_,pt) = me nxt p [] pt; f2str f;
   226 (*[z = 1 / 2, z = -1 / 4]*)
   227 show_pt pt; 
   228 val SOME f = parseNEW ctxt "[z=1/2, z=-1/4]";
   229 *}
   230 
   231 subsection {*partial fraction decomposition*}
   232 subsubsection {*solution of the equation*}
   233 ML {*
   234 val SOME solutions = parseNEW ctxt "[z=1/2, z=-1/4]";
   235 term2str solutions;
   236 atomty solutions;
   237 *}
   238 
   239 subsubsection {*get solutions out of list*}
   240 text {*in isac's CTP-based programming language: let$ $s_1 = NTH 1$ solutions; $s_2 = NTH 2...$*}
   241 ML {*
   242 val Const ("List.list.Cons", _) $ s_1 $ (Const ("List.list.Cons", _) $
   243       s_2 $ Const ("List.list.Nil", _)) = solutions;
   244 term2str s_1;
   245 term2str s_2;
   246 *}
   247 
   248 ML {* (*Solutions as Denominator --> Denominator1 = z - Zeropoint1, Denominator2 = z-Zeropoint2,...*)
   249 val xx = HOLogic.dest_eq s_1;
   250 val s_1' = HOLogic.mk_binop "Groups.minus_class.minus" xx;
   251 val xx = HOLogic.dest_eq s_2;
   252 val s_2' = HOLogic.mk_binop "Groups.minus_class.minus" xx;
   253 term2str s_1';
   254 term2str s_2';
   255 *}
   256 
   257 subsubsection {*build expression*}
   258 text {*in isac's CTP-based programming language: let s_1 = Take numerator / (s_1 * s_2)*}
   259 ML {*
   260 (*The Main Denominator is the multiplikation of the partial fraction denominators*)
   261 val denominator' = HOLogic.mk_binop "Groups.times_class.times" (s_1', s_2') ;
   262 val SOME numerator = parseNEW ctxt "3::real";
   263 
   264 val expr' = HOLogic.mk_binop "Rings.inverse_class.divide" (numerator, denominator');
   265 term2str expr';
   266 *}
   267 
   268 subsubsection {*Ansatz - create partial fractions out of our expression*}
   269 ML {*Context.theory_name thy = "Isac"*}
   270 
   271 axiomatization where
   272   ansatz2: "n / (a*b) = A/a + B/(b::real)" and
   273   multiply_eq2: "(n / (a*b) = A/a + B/b) = (a*b*(n  / (a*b)) = a*b*(A/a + B/b))"
   274 
   275 ML {*
   276 (*we use our ansatz2 to rewrite our expression and get an equilation with our expression on the left and the partial fractions of it on the right side*)
   277 val SOME (t1,_) = rewrite_ @{theory Isac} e_rew_ord e_rls false @{thm ansatz2} expr';
   278 term2str t1; atomty t1;
   279 val eq1 = HOLogic.mk_eq (expr', t1);
   280 term2str eq1;
   281 *}
   282 ML {*
   283 (*eliminate the demoninators by multiplying the left and the right side with the main denominator*)
   284 val SOME (eq2,_) = rewrite_ @{theory Isac} e_rew_ord e_rls false @{thm multiply_eq2} eq1;
   285 term2str eq2;
   286 *}
   287 ML {*
   288 (*simplificatoin*)
   289 val SOME (eq3,_) = rewrite_set_ @{theory Isac} false norm_Rational eq2;
   290 term2str eq3; (*?A ?B not simplified*)
   291 *}
   292 ML {*
   293 val SOME fract1 =
   294   parseNEW ctxt "(z - 1 / 2) * (z - -1 / 4) * (A / (z - 1 / 2) + B / (z - -1 / 4))"; (*A B !*)
   295 val SOME (fract2,_) = rewrite_set_ @{theory Isac} false norm_Rational fract1;
   296 term2str fract2 = "(A + -2 * B + 4 * A * z + 4 * B * z) / 4";
   297 (*term2str fract2 = "A * (1 / 4 + z) + B * (-1 / 2 + z)" would be more traditional*)
   298 *}
   299 ML {*
   300 val (numerator, denominator) = HOLogic.dest_eq eq3;
   301 val eq3' = HOLogic.mk_eq (numerator, fract1); (*A B !*)
   302 term2str eq3';
   303 (*MANDATORY: simplify (and remove denominator) otherwise 3 = 0*)
   304 val SOME (eq3'' ,_) = rewrite_set_ @{theory Isac} false norm_Rational eq3';
   305 term2str eq3'';
   306 *}
   307 ML {*Context.theory_name thy = "Isac"(*==================================================*)*}
   308 
   309 subsubsection {*get first koeffizient*}
   310 
   311 ML {*
   312 (*substitude z with the first zeropoint to get A*)
   313 val SOME (eq4_1,_) = rewrite_terms_ @{theory Isac} e_rew_ord e_rls [s_1] eq3'';
   314 term2str eq4_1;
   315 
   316 val SOME (eq4_2,_) = rewrite_set_ @{theory Isac} false norm_Rational eq4_1;
   317 term2str eq4_2;
   318 
   319 val fmz = ["equality (3 = 3 * A / (4::real))", "solveFor A","solutions L"];
   320 val (dI',pI',mI') =("Isac", ["univariate","equation"], ["no_met"]);
   321 (*solve the simple linear equilation for A TODO: return eq, not list of eq*)
   322 val (p,_,fa,nxt,_,pt) = CalcTreeTEST [(fmz, (dI',pI',mI'))];
   323 val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   324 val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   325 val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   326 val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   327 val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   328 val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   329 val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   330 val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   331 val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   332 val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   333 val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   334 val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   335 val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   336 val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   337 val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   338 val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   339 val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   340 val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   341 val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   342 val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   343 val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   344 val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   345 val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   346 val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   347 val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   348 val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   349 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; 
   350 f2str fa;
   351 *}
   352 
   353 subsubsection {*get second koeffizient*}
   354 ML {*thy*}
   355 
   356 ML {*
   357 (*substitude z with the second zeropoint to get B*)
   358 val SOME (eq4b_1,_) = rewrite_terms_ @{theory Isac} e_rew_ord e_rls [s_2] eq3'';
   359 term2str eq4b_1;
   360 
   361 val SOME (eq4b_2,_) = rewrite_set_ @{theory Isac} false norm_Rational eq4b_1;
   362 term2str eq4b_2;
   363 *}
   364 ML {*
   365 (*solve the simple linear equilation for B TODO: return eq, not list of eq*)
   366 val fmz = ["equality (3 = -3 * B / (4::real))", "solveFor B","solutions L"];
   367 val (dI',pI',mI') =("Isac", ["univariate","equation"], ["no_met"]);
   368 val (p,_,fb,nxt,_,pt) = CalcTreeTEST [(fmz, (dI',pI',mI'))];
   369 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   370 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   371 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   372 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   373 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   374 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   375 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   376 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   377 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   378 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   379 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   380 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   381 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   382 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   383 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   384 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   385 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   386 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   387 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   388 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   389 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   390 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   391 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   392 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   393 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   394 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   395 val (p,_,fb,nxt,_,pt) = me nxt p [] pt; 
   396 f2str fb;
   397 *}
   398 
   399 ML {* (*check koeffizients*)
   400 if f2str fa = "[A = 4]" then () else error "part.fract. eq4_1";
   401 if f2str fb = "[B = -4]" then () else error "part.fract. eq4_1";
   402 *}
   403 
   404 subsubsection {*substitute expression with solutions*}
   405 ML {*
   406 *}
   407 ML {*thy*}
   408 
   409 section {*Implement the Specification and the Method \label{spec-meth}*}
   410 text{*==============================================*}
   411 subsection{*Define the Field Descriptions for the specification*}
   412 consts
   413   filterExpression  :: "bool => una"
   414   stepResponse      :: "bool => una"
   415 
   416 subsection{*Define the Specification*}
   417 ML {*
   418 store_pbt
   419  (prep_pbt thy "pbl_SP" [] e_pblID
   420  (["SignalProcessing"], [], e_rls, NONE, []));
   421 store_pbt
   422  (prep_pbt thy "pbl_SP_Ztrans" [] e_pblID
   423  (["Z_Transform","SignalProcessing"], [], e_rls, NONE, []));
   424 *}
   425 ML {*thy*}
   426 ML {*
   427 store_pbt
   428  (prep_pbt thy "pbl_SP_Ztrans_inv" [] e_pblID
   429  (["inverse", "Z_Transform", "SignalProcessing"],
   430   [("#Given" ,["filterExpression X_eq"]),
   431    ("#Find"  ,["stepResponse n_eq"])
   432   ],
   433   append_rls "e_rls" e_rls [(*for preds in where_*)], NONE, 
   434   [["SignalProcessing","Z_Transform","inverse"]]));
   435 
   436 show_ptyps();
   437 get_pbt ["inverse","Z_Transform","SignalProcessing"];
   438 *}
   439 
   440 subsection {*Define Name and Signature for the Method*}
   441 consts
   442   InverseZTransform :: "[bool, bool] => bool"
   443     ("((Script InverseZTransform (_ =))// (_))" 9)
   444 
   445 subsection {*Setup Parent Nodes in Hierarchy of Method*}
   446 ML {*
   447 store_met
   448  (prep_met thy "met_SP" [] e_metID
   449  (["SignalProcessing"], [],
   450    {rew_ord'="tless_true", rls'= e_rls, calc = [], srls = e_rls, prls = e_rls,
   451     crls = e_rls, nrls = e_rls}, "empty_script"));
   452 store_met
   453  (prep_met thy "met_SP_Ztrans" [] e_metID
   454  (["SignalProcessing", "Z_Transform"], [],
   455    {rew_ord'="tless_true", rls'= e_rls, calc = [], srls = e_rls, prls = e_rls,
   456     crls = e_rls, nrls = e_rls}, "empty_script"));
   457 *}
   458 ML {*
   459 store_met
   460  (prep_met thy "met_SP_Ztrans_inv" [] e_metID
   461  (["SignalProcessing", "Z_Transform", "inverse"], 
   462   [("#Given" ,["filterExpression X_eq"]),
   463    ("#Find"  ,["stepResponse n_eq"])
   464   ],
   465    {rew_ord'="tless_true", rls'= e_rls, calc = [], srls = e_rls, prls = e_rls,
   466     crls = e_rls, nrls = e_rls},
   467   "empty_script"
   468  ));
   469 *}
   470 ML {*
   471 store_met
   472  (prep_met thy "met_SP_Ztrans_inv" [] e_metID
   473  (["SignalProcessing", "Z_Transform", "inverse"], 
   474   [("#Given" ,["filterExpression X_eq"]),
   475    ("#Find"  ,["stepResponse n_eq"])
   476   ],
   477    {rew_ord'="tless_true", rls'= e_rls, calc = [], srls = e_rls, prls = e_rls,
   478     crls = e_rls, nrls = e_rls},
   479   "Script InverseZTransform (Xeq::bool) =" ^
   480   " (let X = Take Xeq;" ^
   481   "      X = Rewrite ruleZY False X" ^
   482   "  in X)"
   483  ));
   484 *}
   485 ML {*
   486 show_mets();
   487 *}
   488 ML {*
   489 get_met ["SignalProcessing","Z_Transform","inverse"];
   490 *}
   491 
   492 section {*Program in CTP-based language \label{prog-steps}*}
   493 text{*=================================*}
   494 subsection {*Stepwise extend Program*}
   495 ML {*
   496 val str = 
   497 "Script InverseZTransform (Xeq::bool) =" ^
   498 " Xeq";
   499 *}
   500 ML {*
   501 val str = 
   502 "Script InverseZTransform (Xeq::bool) =" ^ (*(1/z) instead of z ^^^ -1*)
   503 " (let X = Take Xeq;" ^
   504 "      X' = Rewrite ruleZY False X;" ^ (*z * denominator*)
   505 "      X' = (Rewrite_Set norm_Rational False) X'" ^ (*simplify*)
   506 "  in X)";
   507 (*NONE*)
   508 "Script InverseZTransform (Xeq::bool) =" ^ (*(1/z) instead of z ^^^ -1*)
   509 " (let X = Take Xeq;" ^
   510 "      X' = Rewrite ruleZY False X;" ^ (*z * denominator*)
   511 "      X' = (Rewrite_Set norm_Rational False) X';" ^ (*simplify*)
   512 "      X' = (SubProblem (Isac',[pqFormula,degree_2,polynomial,univariate,equation], [no_met])   " ^
   513     "                 [BOOL e_e, REAL v_v])" ^
   514 "  in X)";
   515 *}
   516 ML {*
   517 val str = 
   518 "Script InverseZTransform (Xeq::bool) =" ^ (*(1/z) instead of z ^^^ -1*)
   519 " (let X = Take Xeq;" ^
   520 "      X' = Rewrite ruleZY False X;" ^ (*z * denominator*)
   521 "      X' = (Rewrite_Set norm_Rational False) X';" ^ (*simplify*)
   522 "      funterm = rhs X'" ^ (*drop X'= for equation solving*)
   523 "  in X)";
   524 *}
   525 ML {*
   526 val str = 
   527 "Script InverseZTransform (X_eq::bool) =" ^ (*(1/z) instead of z ^^^ -1*)
   528 " (let X = Take X_eq;" ^
   529 "      X' = Rewrite ruleZY False X;" ^ (*z * denominator*)
   530 "      X' = (Rewrite_Set norm_Rational False) X';" ^ (*simplify*)
   531 "      (X'_z::real) = lhs X';" ^
   532 "      (z::real) = argument_in X'_z;" ^
   533 "      (funterm::real) = rhs X';" ^ (*drop X' z = for equation solving*)
   534 "      (denom::real) = get_denominator funterm;" ^ (*get_denominator*)
   535 "      (equ::bool) = (denom = (0::real));" ^
   536 "      (L_L::bool list) =                                    " ^
   537 "            (SubProblem (Test',                            " ^
   538 "                         [linear,univariate,equation,test]," ^
   539 "                         [Test,solve_linear])              " ^
   540 "                        [BOOL equ, REAL z])              " ^
   541 "  in X)"
   542 ;
   543 
   544 parse thy str;
   545 val sc = ((inst_abs thy) o term_of o the o (parse thy)) str;
   546 atomty sc;
   547 
   548 *}
   549 
   550 text {*
   551 This ruleset contains all functions that are in the script; 
   552 The evaluation of the functions is done by rewriting using this ruleset.
   553 *}
   554 
   555 ML {*
   556 val srls = Rls {id="srls_InverseZTransform", 
   557 		  preconds = [], rew_ord = ("termlessI",termlessI), 
   558 		  erls = append_rls "erls_in_srls_InverseZTransform" e_rls
   559 				    [(*for asm in NTH_CONS ...*) Calc ("Orderings.ord_class.less",eval_equ "#less_"),
   560 				     (*2nd NTH_CONS pushes n+-1 into asms*) Calc("Groups.plus_class.plus", eval_binop "#add_")
   561 				    ], 
   562   srls = Erls, calc = [],
   563 		  rules =
   564     [Thm ("NTH_CONS",num_str @{thm NTH_CONS}),
   565 			     Calc("Groups.plus_class.plus", eval_binop "#add_"),
   566 			     Thm ("NTH_NIL",num_str @{thm NTH_NIL}),
   567 			     Calc("Tools.lhs", eval_lhs"eval_lhs_"), (*<=== ONLY USED*)
   568 			     Calc("Tools.rhs", eval_rhs"eval_rhs_"), (*<=== ONLY USED*)
   569 			     Calc("Atools.argument'_in", eval_argument_in "Atools.argument'_in"),
   570      Calc("Rational.get_denominator",
   571        eval_get_denominator "Rational.get_denominator")
   572 			    ],
   573 		  scr = EmptyScr};
   574 *}
   575 
   576 
   577 subsection {*Store Final Version of Program for Execution*}
   578 ML {*
   579 store_met
   580  (prep_met thy "met_SP_Ztrans_inv" [] e_metID
   581  (["SignalProcessing", "Z_Transform", "inverse"], 
   582   [("#Given" ,["filterExpression X_eq"]),
   583    ("#Find"  ,["stepResponse n_eq"])
   584   ],
   585    {rew_ord'="tless_true", rls'= e_rls, calc = [], srls = srls, 
   586     prls = e_rls,
   587     crls = e_rls, nrls = e_rls},
   588 "Script InverseZTransform (X_eq::bool) =" ^ (*(1/z) instead of z ^^^ -1*)
   589 " (let X = Take X_eq;" ^
   590 "      X' = Rewrite ruleZY False X;" ^ (*z * denominator*)
   591 "      X' = (Rewrite_Set norm_Rational False) X';" ^ (*simplify*)
   592 "      (X'_z::real) = lhs X';" ^ (**)
   593 "      (zzz::real) = argument_in X'_z;" ^ (**)
   594 "      (funterm::real) = rhs X';" ^ (*drop X' z = for equation solving*)
   595 "      (denom::real) = get_denominator funterm;" ^ (*get_denominator*)
   596 "      (equ::bool) = (denom = (0::real));" ^
   597 
   598 "      (L_L::bool list) = (SubProblem (PolyEq'," ^
   599 "          [abcFormula,degree_2,polynomial,univariate,equation],[no_met])" ^
   600 "        [BOOL equ, REAL zzz])              " ^
   601 "  in X)"
   602  ));
   603 *}
   604 
   605 subsection {*Check the Program*}
   606 
   607 subsubsection {*Check the formalization*}
   608 ML {*
   609 val fmz = ["filterExpression (X  = 3 / (z - 1/4 + -1/8 * (1/(z::real))))", 
   610   "stepResponse (x[n::real]::bool)"];
   611 val (dI,pI,mI) = ("Isac", ["inverse", "Z_Transform", "SignalProcessing"], 
   612   ["SignalProcessing","Z_Transform","inverse"]);
   613 
   614 val ([(1, [1], "#Given", Const ("Inverse_Z_Transform.filterExpression", _),
   615             [Const ("HOL.eq", _) $ _ $ _]),
   616            (2, [1], "#Find", Const ("Inverse_Z_Transform.stepResponse", _),
   617             [Free ("x", _) $ _])],
   618           _) = prep_ori fmz thy ((#ppc o get_pbt) pI);
   619 *}
   620 ML {*
   621 val Script sc = (#scr o get_met) ["SignalProcessing","Z_Transform","inverse"];
   622 atomty sc;
   623 *}
   624 
   625 subsubsection {*Stepwise check the program*}
   626 ML {*
   627 trace_rewrite := false;
   628 trace_script := false; print_depth 9;
   629 val fmz = ["filterExpression (X z = 3 / (z - 1/4 + -1/8 * (1/(z::real))))", 
   630   "stepResponse (x[n::real]::bool)"];
   631 val (dI,pI,mI) = ("Isac", ["inverse", "Z_Transform", "SignalProcessing"], 
   632   ["SignalProcessing","Z_Transform","inverse"]);
   633 val (p,_,f,nxt,_,pt)  = CalcTreeTEST [(fmz, (dI,pI,mI))];
   634 *}
   635 ML {*
   636 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "Add_Given";
   637 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "Add_Find";
   638 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "Specify_Theory";
   639 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "Specify_Problem";
   640 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "Specify_Method";
   641 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "nxt = Apply_Method";
   642 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "nxt = Rewrite (ruleZY, Inverse_Z_Transform.ruleZY) --> X z = 3 / (z - 1 / 4 + -1 / 8 * (1 / z))"; (*TODO naming!*)
   643 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "nxt = Rewrite_Set norm_Rational --> X' z = 3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z)))";
   644 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "nxt = SubProblem";
   645 *}
   646 text {* Instead of nxt = Subproblem above we had Empty_Tac; the search for the reason 
   647   considered the following points:
   648   # what shows show_pt pt; ...
   649     (([2], Res), ?X' z = 24 / (-1 + -2 * z + 8 * z ^^^ 2))] ..calculation ok,
   650     but no "next" step found: should be "nxt = Subproblem" ?!?
   651   # what shows trace_script := true; we read bottom up ...
   652     @@@ next   leaf 'SubProbfrom
   653      (PolyEq', [abcFormula, degree_2, polynomial, univariate, equation],
   654       no_meth)
   655      [BOOL equ, REAL z]' ---> STac 'SubProblem
   656      (PolyEq', [abcFormula, degree_2, polynomial, univariate, equation],
   657       no_meth)
   658      [BOOL (-1 + -2 * z + 8 * z ^^^ 2 = 0), REAL z]'
   659     ... and see the SubProblem with correct arguments from searching next step
   660     (program text !!!--->!!! STac (script tactic) with arguments evaluated.)
   661   # do we have the right Script ...difference in the argumentsdifference in the arguments
   662     val Script s = (#scr o get_met) ["SignalProcessing","Z_Transform","inverse"];
   663     writeln (term2str s);
   664     ... shows the right script.difference in the arguments
   665   # test --- why helpless here ? --- shows: replace no_meth by [no_meth] in Script
   666 *}
   667 
   668 ML {*
   669 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "nxt = Model_Problem";
   670 *}
   671 text {* Instead of nxt = Model_Problem above we had Empty_Tac; the search for the reason 
   672   considered the following points:difference in the arguments
   673   # comparison with subsection { *solve equation* }: there solving this equation works,
   674     so there must be some difference in the arguments of the Subproblem:
   675     RIGHT: we had [no_meth] instead of [no_met] ;-))
   676 *}
   677 ML {*
   678 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "nxt = Add_Given equality (-1 + -2 * z + 8 * z ^^^ 2 = 0)";
   679 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "nxt = Add_Given solveFor z";
   680 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "nxt = Add_Find solutions z_i";
   681 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "nxt = Specify_Theory Isac";
   682 *}
   683 text {* We had "nxt = Empty_Tac instead Specify_Theory; 
   684   the search for the reason considered the following points:
   685   # was there an error message ? NO --ok
   686   # has "nxt = Add_Find" been inserted in pt: get_obj g_pbl pt (fst p); YES --ok
   687   # what is the returned "formula": print_depth 999; f; print_depth 999; --
   688     {Find = [Correct "solutions z_i"], With = [], 
   689      Given = [Correct "equality (-1 + -2 * z + 8 * z ^^^ 2 = 0)", Correct "solveFor z"],
   690      Where = [False "matches (z = 0) (-1 + -2 * z + 8 * z ^^^ 2 = 0) |\n
   691                      matches (?b * z = 0) (-1 + -2 * z + 8 * z ^^^ 2 = 0) |\n
   692                      matches (?a + z = 0) (-1 + -2 * z + 8 * z ^^^ 2 = 0) |\n
   693                      matches (?a + ?b * z = 0) (-1 + -2 * z + 8 * z ^^^ 2 = 0)"],
   694      Relate = []}
   695      -- the only False is the reason: the Where (the precondition) is False for good reasons:
   696      the precondition seems to check for linear equations, not for the one we want to solve!
   697   Removed this error by correcting the Script
   698   from SubProblem (PolyEq', [linear,univariate,equation,test], [Test,solve_linear]
   699   to   SubProblem (PolyEq', [abcFormula,degree_2,polynomial,univariate,equation],
   700                    [PolyEq,solve_d2_polyeq_abc_equation]
   701   You find the appropriate type of equation at
   702     http://www.ist.tugraz.at/projects/isac/www/kbase/pbl/index_pbl.html
   703   and the respective method in Knowledge/PolyEq.thy at the respective store_pbt.
   704   Or you leave the selection of the appropriate type to isac as done in the final Script ;-))
   705 *}
   706 ML {*
   707 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "nxt = Specify_Problem [abcFormula,...";
   708 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "nxt = Specify_Method [PolyEq,solve_d2_polyeq_abc_equation";
   709 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "nxt = Apply_Method [PolyEq,solve_d2_polyeq_abc_equation";
   710 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "nxt = Rewrite_Set_Inst ([(bdv, z)], d2_polyeq_abcFormula_simplify";
   711 show_pt pt;
   712 *}
   713 ML {*
   714 *}
   715 ML {*
   716 *}
   717 ML {*
   718 *}
   719 
   720 
   721 section {*Write Tests for Crucial Details*}
   722 text{*===================================*}
   723 ML {*
   724 
   725 *}
   726 
   727 section {*Integrate Program into Knowledge*}
   728 ML {*
   729 @{theory Isac}
   730 *}
   731 
   732 end
   733