test/Tools/isac/ADDTESTS/course/SignalProcess/Build_Inverse_Z_Transform.thy
author Jan Rocnik <jan.rocnik@student.tugraz.at>
Sun, 22 Jan 2012 21:42:56 +0100
changeset 42364 f82f2b7fef61
parent 42363 82d438b5c597
child 42365 872589b4852c
permissions -rwxr-xr-x
Now getting BOTH numerators and small bugfix in Ansatz
     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 
   147 subsubsection {*get the denominator and numerator out of a fraction*}
   148 text {*get denominator should become a constant for the isabelle parser: *}
   149 
   150 consts
   151   get_denominator :: "real => real"
   152   get_numerator :: "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 Rational.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 text {*get numerator should also become a constant for the isabelle parser: *}
   176 
   177 ML {*
   178 fun eval_get_numerator (thmid:string) _ 
   179 		      (t as Const ("Rational.get_numerator", _) $
   180               (Const ("Rings.inverse_class.divide", _) $num
   181                 $denom )) thy = 
   182         SOME (mk_thmid thmid "" 
   183             (Print_Mode.setmp [] (Syntax.string_of_term (thy2ctxt thy)) num) "", 
   184 	          Trueprop $ (mk_equality (t, num)))
   185   | eval_get_numerator _ _ _ _ = NONE; 
   186 *}
   187 
   188 text {*
   189 We discovered severell problems by implementing the get_numerator function.
   190 Remember when putting new functions to Isac, put them in a thy file and rebuild 
   191 isac, also put them in the ruleset for the script!
   192 *}
   193 
   194 subsection {*solve equation*}
   195 text {*this type of equation if too general for the present program*}
   196 ML {*
   197 "----------- Minisubplb/100-init-rootp (*OK*)bl.sml ---------------------";
   198 val denominator = parseNEW ctxt "z^^^2 - 1/4*z - 1/8 = 0";
   199 val fmz = ["equality (z^^^2 - 1/4*z - 1/8 = (0::real))", "solveFor z","solutions L"];
   200 val (dI',pI',mI') =("Isac", ["univariate","equation"], ["no_met"]);
   201 (*                           ^^^^^^^^^^^^^^^^^^^^^^ TODO: ISAC determines type of eq*)
   202 *}
   203 text {*Does the Equation Match the Specification ?*}
   204 ML {*
   205 match_pbl fmz (get_pbt ["univariate","equation"]);
   206 *}
   207 ML {*Context.theory_name thy = "Isac"(*==================================================*)*}
   208 
   209 ML {*
   210 val denominator = parseNEW ctxt "-1 + -2 * z + 8 * z ^^^ 2 = 0";
   211 val fmz =                                            (*specification*)
   212   ["equality (-1 + -2 * z + 8 * z ^^^ 2 = (0::real))", (*equality*)
   213    "solveFor z",                                     (*bound variable*)
   214    "solutions L"];                                   (*identifier for solution*)
   215 
   216 val (dI',pI',mI') =
   217   ("Isac", ["abcFormula","degree_2","polynomial","univariate","equation"], ["no_met"]);
   218 *}
   219 text {*Does the Other Equation Match the Specification ?*}
   220 ML {*
   221 match_pbl fmz (get_pbt ["abcFormula","degree_2","polynomial","univariate","equation"]);
   222 *}
   223 text {*Solve Equation Stepwise*}
   224 ML {*
   225 *}
   226 ML {*
   227 val (p,_,f,nxt,_,pt) = CalcTreeTEST [(fmz, (dI',pI',mI'))];
   228 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   229 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   230 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   231 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   232 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   233 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   234 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   235 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   236 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   237 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   238 val (p,_,f,nxt,_,pt) = me nxt p [] pt;         
   239 val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt =..,Check_elementwise "Assumptions")*)
   240 val (p,_,f,nxt,_,pt) = me nxt p [] pt;         
   241 val (p,_,f,nxt,_,pt) = me nxt p [] pt; f2str f;
   242 (*[z = 1 / 2, z = -1 / 4]*)
   243 show_pt pt; 
   244 val SOME f = parseNEW ctxt "[z=1/2, z=-1/4]";
   245 *}
   246 
   247 subsection {*partial fraction decomposition*}
   248 subsubsection {*solution of the equation*}
   249 ML {*
   250 val SOME solutions = parseNEW ctxt "[z=1/2, z=-1/4]";
   251 term2str solutions;
   252 atomty solutions;
   253 *}
   254 
   255 subsubsection {*get solutions out of list*}
   256 text {*in isac's CTP-based programming language: let$ $s_1 = NTH 1$ solutions; $s_2 = NTH 2...$*}
   257 ML {*
   258 val Const ("List.list.Cons", _) $ s_1 $ (Const ("List.list.Cons", _) $
   259       s_2 $ Const ("List.list.Nil", _)) = solutions;
   260 term2str s_1;
   261 term2str s_2;
   262 *}
   263 
   264 ML {* (*Solutions as Denominator --> Denominator1 = z - Zeropoint1, Denominator2 = z-Zeropoint2,...*)
   265 val xx = HOLogic.dest_eq s_1;
   266 val s_1' = HOLogic.mk_binop "Groups.minus_class.minus" xx;
   267 val xx = HOLogic.dest_eq s_2;
   268 val s_2' = HOLogic.mk_binop "Groups.minus_class.minus" xx;
   269 term2str s_1';
   270 term2str s_2';
   271 *}
   272 text {* for the programming language a function 
   273   collecting all the above manipulations is helpful*}
   274 ML {*
   275 fun mk_minus_1 T = Free("-1", T); (*TODO DELETE WITH numbers_to_string*)
   276 fun flip_sign t = (*TODO improve for use in factors_from_solution: -(-1) etc*)
   277   let val minus_1 = t |> type_of |> mk_minus_1
   278   in HOLogic.mk_binop "Groups.times_class.times" (minus_1, t) end;
   279 fun fac_from_sol s =
   280   let val (lhs, rhs) = HOLogic.dest_eq s
   281   in HOLogic.mk_binop "Groups.plus_class.plus" (lhs, flip_sign rhs) end;
   282 *}
   283 ML {*
   284 e_term
   285 *}
   286 ML {*
   287 fun mk_prod prod [] =
   288       if prod = e_term then error "mk_prod called with []" else prod
   289   | mk_prod prod (t :: []) =
   290       if prod = e_term then t else HOLogic.mk_binop "Groups.times_class.times" (prod, t)
   291   | mk_prod prod (t1 :: t2 :: ts) =
   292         if prod = e_term 
   293         then 
   294            let val p = HOLogic.mk_binop "Groups.times_class.times" (t1, t2)
   295            in mk_prod p ts end 
   296         else 
   297            let val p = HOLogic.mk_binop "Groups.times_class.times" (prod, t1)
   298            in mk_prod p (t2 :: ts) end 
   299 *}
   300 ML {*
   301 *}
   302 ML {*
   303 (*probably keept these test in test/Tools/isac/...
   304 (*mk_prod e_term [];*)
   305 
   306 val prod = mk_prod e_term [str2term "x + 123"]; 
   307 term2str prod = "x + 123";
   308 
   309 val sol = str2term "[z = 1 / 2, z = -1 / 4]";
   310 val sols = HOLogic.dest_list sol;
   311 val facs = map fac_from_sol sols;
   312 val prod = mk_prod e_term facs; 
   313 term2str prod = "(z + -1 * (1 / 2)) * (z + -1 * (-1 / 4))";
   314 
   315 val prod = mk_prod e_term [str2term "x + 1", str2term "x + 2", str2term "x + 3"]; 
   316 term2str prod = "(x + 1) * (x + 2) * (x + 3)";
   317 *)
   318 
   319 fun factors_from_solution sol = 
   320   let val ts = HOLogic.dest_list sol
   321   in mk_prod e_term (map fac_from_sol ts) end;
   322 (*
   323 val sol = str2term "[z = 1 / 2, z = -1 / 4]";
   324 val fs = factors_from_solution sol;
   325 term2str fs = "(z + -1 * (1 / 2)) * (z + -1 * (-1 / 4))"
   326 *)
   327 *}
   328 text {* This function needs to be packed such that it can be evaluated by the Lucas-Interpreter:
   329   # shift these functions into the related Equation.thy
   330   #  -- compare steps done with get_denominator above
   331   # done 02.12.2011 moved to PartialFractions.thy
   332   *}
   333 ML {*
   334 (*("factors_from_solution", ("Partial_Fractions.factors_from_solution", eval_factors_from_solution ""))*)
   335 fun eval_factors_from_solution (thmid:string) _
   336      (t as Const ("Partial_Fractions.factors_from_solution", _) $ sol) thy =
   337        ((let val prod = factors_from_solution sol
   338          in SOME (mk_thmid thmid ""
   339            (Print_Mode.setmp [] (Syntax.string_of_term (thy2ctxt thy)) prod) "",
   340                Trueprop $ (mk_equality (t, prod)))
   341          end)
   342        handle _ => NONE)
   343  | eval_factors_from_solution _ _ _ _ = NONE;
   344 *}
   345 
   346 text {*
   347 The tracing output of the calc tree after apllying this function was
   348 24 / factors_from_solution [z = 1/ 2, z = -1 / 4])]  and the next step
   349  val nxt = ("Empty_Tac", ...): tac'_).
   350 These observations indicate, that the Lucas-Interpreter (LIP) does 
   351 not know how to evaluate factors_from_solution, so there is something 
   352 wrong or missing.
   353 
   354 # First we isolate the difficulty in the program as follows:
   355   :
   356   "      (L_L::bool list) = (SubProblem (PolyEq'," ^
   357   "          [abcFormula,degree_2,polynomial,univariate,equation],[no_met])" ^
   358   "        [BOOL equ, REAL zzz]);              " ^
   359   "      (facs::real) = factors_from_solution L_L;" ^
   360   "      (foo::real) = Take facs" ^
   361   :
   362 and see
   363   [
   364   (([], Frm), Problem (Isac, [inverse, Z_Transform, SignalProcessing])),
   365   (([1], Frm), X z = 3 / (z - 1 / 4 + -1 / 8 * (1 / z))),
   366   (([1], Res), ?X' z = 3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z)))),
   367   (([2], Res), ?X' z = 24 / (-1 + -2 * z + 8 * z ^^^ 2)),
   368   (([3], Pbl), solve (-1 + -2 * z + 8 * z ^^^ 2 = 0, z)),
   369   (([3,1], Frm), -1 + -2 * z + 8 * z ^^^ 2 = 0),
   370   (([3,1], Res), z = (- -2 + sqrt (-2 ^^^ 2 - 4 * 8 * -1)) / (2 * 8) |
   371                  z = (- -2 - sqrt (-2 ^^^ 2 - 4 * 8 * -1)) / (2 * 8)),
   372   (([3,2], Res), z = 1 / 2 | z = -1 / 4),
   373   (([3,3], Res), [z = 1 / 2, z = -1 / 4]),
   374   (([3,4], Res), [z = 1 / 2, z = -1 / 4]),
   375   (([3], Res), [z = 1 / 2, z = -1 / 4]),
   376   (([4], Frm), factors_from_solution [z = 1 / 2, z = -1 / 4])]
   377 in particular that
   378   (([3], Pbl), solve (-1 + -2 * z + 8 * z ^^^ 2 = 0, z)),
   379 shows the equation which has been created in the program by
   380   "      (denom::real) = get_denominator funterm;" ^ (*get_denominator*)
   381   "      (equ::bool) = (denom = (0::real));" ^
   382 # 'get_denominator' has been evaluated successfully, but not factors_from_solution.
   383 So we stepwise compare with an analogous case, get_denominator 
   384 successfully done above: We know that LIP evaluates expressions in the 
   385 program by use of the "srls", so we
   386 # try to get the original srls
   387 
   388   val {srls, ...} = get_met ["SignalProcessing","Z_Transform","inverse"];
   389 
   390 # create 2 good example terms
   391   val SOME t1 = parseNEW ctxt "get_denominator ((111::real) / 222)";
   392   val SOME t2 = parseNEW ctxt "factors_from_solution [(z::real) = 1 / 2, z = -1 / 4]";
   393 
   394 # rewrite the terms using srls
   395   rewrite_set_ thy true srls t1;
   396   rewrite_set_ thy true srls t2;
   397 
   398 and we see a difference: t1 gives SOME, t2 gives NONE.
   399 Now we look at the srls:
   400   val srls = Rls {id="srls_InverseZTransform",
   401     :
   402     rules =
   403       [
   404         :
   405         Calc("Rational.get_numerator",
   406           eval_get_numerator "Rational.get_numerator"),
   407         Calc("Partial_Fractions.factors_from_solution",
   408           eval_factors_from_solution "Partial_Fractions.factors_from_solution")
   409       ],
   410     :
   411 
   412 Here everthing is perfect. So the error can only be in the SML code of eval_factors_from_solution.
   413 We try to check the code with an existing test; since the code is in
   414 
   415   src/Tools/isac/Knowledge/Partial_Fractions.thy
   416 
   417 the test should be in
   418 
   419   test/Tools/isac/Knowledge/partial_fractions.sml
   420 
   421 -------------------------------------------------------------------------------
   422 After updating the function get_factors_from solution to a new version and 
   423 putting a testcase to Partial_Fractions.sml we tried again to evaluate the
   424 term with the same result.
   425 We opened the test Test_Isac.thy and saw that everything is working fine.
   426 Also we checked that the test partial_fractions.sml is used in Test_Isac.thy 
   427 
   428 -->  use "Knowledge/partial_fractions.sml"
   429 
   430 and Partial_Fractions.thy is part is part of isac by evaluating
   431 
   432 val thy = @{theory Isac};
   433 
   434 after rebuilding isac again it worked
   435 
   436 *}
   437 
   438 subsubsection {*build expression*}
   439 text {*in isac's CTP-based programming language: let s_1 = Take numerator / (s_1 * s_2)*}
   440 ML {*
   441 (*The Main Denominator is the multiplikation of the partial fraction denominators*)
   442 val denominator' = HOLogic.mk_binop "Groups.times_class.times" (s_1', s_2') ;
   443 val SOME numerator = parseNEW ctxt "3::real";
   444 
   445 val expr' = HOLogic.mk_binop "Rings.inverse_class.divide" (numerator, denominator');
   446 term2str expr';
   447 *}
   448 
   449 subsubsection {*Ansatz - create partial fractions out of our expression*}
   450 ML {*Context.theory_name thy = "Isac"*}
   451 
   452 axiomatization where
   453   ansatz2: "n / (a*b) = A/a + B/(b::real)" and
   454   multiply_eq2: "((n::real) / (a*b) = A/a + B/b) = (a*b*(n  / (a*b)) = a*b*(A/a + B/b::real))"  
   455 
   456 ML {*
   457 (*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*)
   458 val SOME (t1,_) = rewrite_ @{theory Isac} e_rew_ord e_rls false @{thm ansatz2} expr';
   459 term2str t1; atomty t1;
   460 val eq1 = HOLogic.mk_eq (expr', t1);
   461 term2str eq1;
   462 *}
   463 ML {*
   464 (*eliminate the demoninators by multiplying the left and the right side with the main denominator*)
   465 val SOME (eq2,_) = rewrite_ @{theory Isac} e_rew_ord e_rls false @{thm multiply_eq2} eq1;
   466 term2str eq2;
   467 *}
   468 ML {*
   469 (*simplificatoin*)
   470 val SOME (eq3,_) = rewrite_set_ @{theory Isac} false norm_Rational eq2;
   471 term2str eq3; (*?A ?B not simplified*)
   472 *}
   473 ML {*
   474 val SOME fract1 =
   475   parseNEW ctxt "(z - 1 / 2) * (z - -1 / 4) * (A / (z - 1 / 2) + B / (z - -1 / 4))"; (*A B !*)
   476 val SOME (fract2,_) = rewrite_set_ @{theory Isac} false norm_Rational fract1;
   477 term2str fract2 = "(A + -2 * B + 4 * A * z + 4 * B * z) / 4";
   478 (*term2str fract2 = "A * (1 / 4 + z) + B * (-1 / 2 + z)" would be more traditional*)
   479 *}
   480 ML {*
   481 val (numerator, denominator) = HOLogic.dest_eq eq3;
   482 val eq3' = HOLogic.mk_eq (numerator, fract1); (*A B !*)
   483 term2str eq3';
   484 (*MANDATORY: simplify (and remove denominator) otherwise 3 = 0*)
   485 val SOME (eq3'' ,_) = rewrite_set_ @{theory Isac} false norm_Rational eq3';
   486 term2str eq3'';
   487 *}
   488 ML {*Context.theory_name thy = "Isac"(*==================================================*)*}
   489 
   490 subsubsection {*Build a rule-set for ansatz*}
   491 text {* the "ansatz" rules violate the principle that each variable on 
   492   the right-hand-side must also occur on the left-hand-side of the rule:
   493   A, B, etc don't.
   494   Thus the rewriter marks these variables with question marks: ?A, ?B, etc.
   495   These question marks can be dropped by "fun drop_questionmarks".
   496   *}
   497 ML {*
   498 val ansatz_rls = prep_rls(
   499   Rls {id = "ansatz_rls", preconds = [], rew_ord = ("dummy_ord",dummy_ord), 
   500 	  erls = e_rls, srls = Erls, calc = [],
   501 	  rules = 
   502 	   [Thm ("ansatz2",num_str @{thm ansatz2}),
   503 	    Thm ("multiply_eq2",num_str @{thm multiply_eq2})
   504 	   ], 
   505 	 scr = EmptyScr});
   506 *}
   507 ML {*
   508 val SOME (ttttt,_) = rewrite_set_ @{theory Isac} false ansatz_rls expr';
   509 *}
   510 ML {*
   511 term2str expr' = "3 / ((z - 1 / 2) * (z - -1 / 4))";
   512 term2str ttttt = "?A / (z - 1 / 2) + ?B / (z - -1 / 4)";
   513 *}
   514 
   515 
   516 subsubsection {*get first koeffizient*}
   517 
   518 ML {*
   519 (*substitude z with the first zeropoint to get A*)
   520 val SOME (eq4_1,_) = rewrite_terms_ @{theory Isac} e_rew_ord e_rls [s_1] eq3'';
   521 term2str eq4_1;
   522 
   523 val SOME (eq4_2,_) = rewrite_set_ @{theory Isac} false norm_Rational eq4_1;
   524 term2str eq4_2;
   525 
   526 val fmz = ["equality (3 = 3 * A / (4::real))", "solveFor A","solutions L"];
   527 val (dI',pI',mI') =("Isac", ["univariate","equation"], ["no_met"]);
   528 (*solve the simple linear equilation for A TODO: return eq, not list of eq*)
   529 val (p,_,fa,nxt,_,pt) = CalcTreeTEST [(fmz, (dI',pI',mI'))];
   530 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Add_Given "equality (3 = 3 * A / 4)"*)
   531 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (* Add_Given "solveFor A"*)
   532 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Add_Find "solutions L"*)
   533 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Specify_Theory "Isac"*)
   534 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Specify_Problem ["normalize", "polynomial", 
   535                                           "univariate", "equation"])*)
   536 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (* Specify_Method ["PolyEq", "normalize_poly"]*)
   537 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Apply_Method ["PolyEq", "normalize_poly"]*)
   538 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Rewrite ("all_left", "PolyEq.all_left")*)
   539 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Rewrite_Set_Inst (["(bdv, A)"], "make_ratpoly_in")*)
   540 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Rewrite_Set "polyeq_simplify"*)
   541 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (**)
   542 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (**)
   543 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Add_Given "equality (3 + -3 / 4 * A = 0)"*)
   544 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Add_Given "solveFor A"*)
   545 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Add_Find "solutions A_i"*)
   546 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (**)
   547 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (**)
   548 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (**)
   549 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Apply_Method ["PolyEq", "solve_d1_polyeq_equation"]*)
   550 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Rewrite_Set_Inst (["(bdv, A)"], "d1_polyeq_simplify")*)
   551 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Rewrite_Set "polyeq_simplify"*)
   552 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Rewrite_Set "norm_Rational_parenthesized"*)
   553 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Or_to_List*)
   554 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Check_elementwise "Assumptions"*)
   555 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Check_Postcond ["degree_1", "polynomial", 
   556                                           "univariate", "equation"]*)
   557 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Check_Postcond ["normalize", "polynomial", 
   558                                           "univariate", "equation"]*)
   559 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*End_Proof'*)
   560 f2str fa;
   561 *}
   562 
   563 subsubsection {*get second koeffizient*}
   564 ML {*thy*}
   565 
   566 ML {*
   567 (*substitude z with the second zeropoint to get B*)
   568 val SOME (eq4b_1,_) = rewrite_terms_ @{theory Isac} e_rew_ord e_rls [s_2] eq3'';
   569 term2str eq4b_1;
   570 
   571 val SOME (eq4b_2,_) = rewrite_set_ @{theory Isac} false norm_Rational eq4b_1;
   572 term2str eq4b_2;
   573 *}
   574 ML {*
   575 (*solve the simple linear equilation for B TODO: return eq, not list of eq*)
   576 val fmz = ["equality (3 = -3 * B / (4::real))", "solveFor B","solutions L"];
   577 val (dI',pI',mI') =("Isac", ["univariate","equation"], ["no_met"]);
   578 val (p,_,fb,nxt,_,pt) = CalcTreeTEST [(fmz, (dI',pI',mI'))];
   579 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   580 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   581 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   582 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   583 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   584 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   585 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   586 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   587 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   588 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   589 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   590 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   591 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   592 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   593 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   594 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   595 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   596 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   597 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   598 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   599 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   600 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   601 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   602 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   603 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   604 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   605 val (p,_,fb,nxt,_,pt) = me nxt p [] pt; 
   606 f2str fb;
   607 *}
   608 
   609 ML {* (*check koeffizients*)
   610 if f2str fa = "[A = 4]" then () else error "part.fract. eq4_1";
   611 if f2str fb = "[B = -4]" then () else error "part.fract. eq4_1";
   612 *}
   613 
   614 subsubsection {*substitute expression with solutions*}
   615 ML {*
   616 *}
   617 ML {*thy*}
   618 
   619 section {*Implement the Specification and the Method \label{spec-meth}*}
   620 text{*==============================================*}
   621 subsection{*Define the Field Descriptions for the specification*}
   622 consts
   623   filterExpression  :: "bool => una"
   624   stepResponse      :: "bool => una"
   625 
   626 subsection{*Define the Specification*}
   627 ML {*
   628 store_pbt
   629  (prep_pbt thy "pbl_SP" [] e_pblID
   630  (["SignalProcessing"], [], e_rls, NONE, []));
   631 store_pbt
   632  (prep_pbt thy "pbl_SP_Ztrans" [] e_pblID
   633  (["Z_Transform","SignalProcessing"], [], e_rls, NONE, []));
   634 *}
   635 ML {*thy*}
   636 ML {*
   637 store_pbt
   638  (prep_pbt thy "pbl_SP_Ztrans_inv" [] e_pblID
   639  (["inverse", "Z_Transform", "SignalProcessing"],
   640   [("#Given" ,["filterExpression X_eq"]),
   641    ("#Find"  ,["stepResponse n_eq"])
   642   ],
   643   append_rls "e_rls" e_rls [(*for preds in where_*)], NONE, 
   644   [["SignalProcessing","Z_Transform","inverse"]]));
   645 
   646 show_ptyps();
   647 get_pbt ["inverse","Z_Transform","SignalProcessing"];
   648 *}
   649 
   650 subsection {*Define Name and Signature for the Method*}
   651 consts
   652   InverseZTransform :: "[bool, bool] => bool"
   653     ("((Script InverseZTransform (_ =))// (_))" 9)
   654 
   655 subsection {*Setup Parent Nodes in Hierarchy of Method*}
   656 ML {*
   657 store_met
   658  (prep_met thy "met_SP" [] e_metID
   659  (["SignalProcessing"], [],
   660    {rew_ord'="tless_true", rls'= e_rls, calc = [], srls = e_rls, prls = e_rls,
   661     crls = e_rls, nrls = e_rls}, "empty_script"));
   662 store_met
   663  (prep_met thy "met_SP_Ztrans" [] e_metID
   664  (["SignalProcessing", "Z_Transform"], [],
   665    {rew_ord'="tless_true", rls'= e_rls, calc = [], srls = e_rls, prls = e_rls,
   666     crls = e_rls, nrls = e_rls}, "empty_script"));
   667 *}
   668 ML {*
   669 store_met
   670  (prep_met thy "met_SP_Ztrans_inv" [] e_metID
   671  (["SignalProcessing", "Z_Transform", "inverse"], 
   672   [("#Given" ,["filterExpression X_eq"]),
   673    ("#Find"  ,["stepResponse n_eq"])
   674   ],
   675    {rew_ord'="tless_true", rls'= e_rls, calc = [], srls = e_rls, prls = e_rls,
   676     crls = e_rls, nrls = e_rls},
   677   "empty_script"
   678  ));
   679 *}
   680 ML {*
   681 store_met
   682  (prep_met thy "met_SP_Ztrans_inv" [] e_metID
   683  (["SignalProcessing", "Z_Transform", "inverse"], 
   684   [("#Given" ,["filterExpression X_eq"]),
   685    ("#Find"  ,["stepResponse n_eq"])
   686   ],
   687    {rew_ord'="tless_true", rls'= e_rls, calc = [], srls = e_rls, prls = e_rls,
   688     crls = e_rls, nrls = e_rls},
   689   "Script InverseZTransform (Xeq::bool) =" ^
   690   " (let X = Take Xeq;" ^
   691   "      X = Rewrite ruleZY False X" ^
   692   "  in X)"
   693  ));
   694 *}
   695 ML {*
   696 show_mets();
   697 *}
   698 ML {*
   699 get_met ["SignalProcessing","Z_Transform","inverse"];
   700 *}
   701 
   702 section {*Program in CTP-based language \label{prog-steps}*}
   703 text{*=================================*}
   704 subsection {*Stepwise extend Program*}
   705 ML {*
   706 val str = 
   707 "Script InverseZTransform (Xeq::bool) =" ^
   708 " Xeq";
   709 *}
   710 ML {*
   711 val str = 
   712 "Script InverseZTransform (Xeq::bool) =" ^ (*(1/z) instead of z ^^^ -1*)
   713 " (let X = Take Xeq;" ^
   714 "      X' = Rewrite ruleZY False X;" ^ (*z * denominator*)
   715 "      X' = (Rewrite_Set norm_Rational False) X'" ^ (*simplify*)
   716 "  in X)";
   717 (*NONE*)
   718 "Script InverseZTransform (Xeq::bool) =" ^ (*(1/z) instead of z ^^^ -1*)
   719 " (let X = Take Xeq;" ^
   720 "      X' = Rewrite ruleZY False X;" ^ (*z * denominator*)
   721 "      X' = (Rewrite_Set norm_Rational False) X';" ^ (*simplify*)
   722 "      X' = (SubProblem (Isac',[pqFormula,degree_2,polynomial,univariate,equation], [no_met])   " ^
   723     "                 [BOOL e_e, REAL v_v])" ^
   724 "  in X)";
   725 *}
   726 ML {*
   727 val str = 
   728 "Script InverseZTransform (Xeq::bool) =" ^ (*(1/z) instead of z ^^^ -1*)
   729 " (let X = Take Xeq;" ^
   730 "      X' = Rewrite ruleZY False X;" ^ (*z * denominator*)
   731 "      X' = (Rewrite_Set norm_Rational False) X';" ^ (*simplify*)
   732 "      funterm = rhs X'" ^ (*drop X'= for equation solving*)
   733 "  in X)";
   734 *}
   735 ML {*
   736 val str = 
   737 "Script InverseZTransform (X_eq::bool) =" ^ (*(1/z) instead of z ^^^ -1*)
   738 " (let X = Take X_eq;" ^
   739 "      X' = Rewrite ruleZY False X;" ^ (*z * denominator*)
   740 "      X' = (Rewrite_Set norm_Rational False) X';" ^ (*simplify*)
   741 "      (X'_z::real) = lhs X';" ^
   742 "      (z::real) = argument_in X'_z;" ^
   743 "      (funterm::real) = rhs X';" ^ (*drop X' z = for equation solving*)
   744 "      (denom::real) = get_denominator funterm;" ^ (*get_denominator*)
   745 "      (equ::bool) = (denom = (0::real));" ^
   746 "      (L_L::bool list) =                                    " ^
   747 "            (SubProblem (Test',                            " ^
   748 "                         [linear,univariate,equation,test]," ^
   749 "                         [Test,solve_linear])              " ^
   750 "                        [BOOL equ, REAL z])              " ^
   751 "  in X)"
   752 ;
   753 
   754 parse thy str;
   755 val sc = ((inst_abs thy) o term_of o the o (parse thy)) str;
   756 atomty sc;
   757 
   758 *}
   759 
   760 text {*
   761 This ruleset contains all functions that are in the script; 
   762 The evaluation of the functions is done by rewriting using this ruleset.
   763 *}
   764 
   765 ML {*
   766 val srls = Rls {id="srls_InverseZTransform", 
   767 		  preconds = [], rew_ord = ("termlessI",termlessI), 
   768 		  erls = append_rls "erls_in_srls_InverseZTransform" e_rls
   769 				    [(*for asm in NTH_CONS ...*) Calc ("Orderings.ord_class.less",eval_equ "#less_"),
   770 				     (*2nd NTH_CONS pushes n+-1 into asms*) Calc("Groups.plus_class.plus", eval_binop "#add_")
   771 				    ], 
   772   srls = Erls, calc = [],
   773 		  rules =
   774     [Thm ("NTH_CONS",num_str @{thm NTH_CONS}),
   775 			     Calc("Groups.plus_class.plus", eval_binop "#add_"),
   776 			     Thm ("NTH_NIL",num_str @{thm NTH_NIL}),
   777 			     Calc("Tools.lhs", eval_lhs"eval_lhs_"), (*<=== ONLY USED*)
   778 			     Calc("Tools.rhs", eval_rhs"eval_rhs_"), (*<=== ONLY USED*)
   779 			     Calc("Atools.argument'_in", eval_argument_in "Atools.argument'_in"),
   780      Calc("Rational.get_denominator", eval_get_denominator "#get_denominator"),
   781      Calc("Rational.get_numerator", eval_get_numerator "#get_numerator"),
   782      Calc("Partial_Fractions.factors_from_solution",
   783        eval_factors_from_solution "#factors_from_solution"),
   784      Calc("Partial_Fractions.drop_questionmarks", eval_drop_questionmarks "#drop_?")
   785 			    ],
   786 		  scr = EmptyScr};
   787 *}
   788 
   789 
   790 subsection {*Store Final Version of Program for Execution*}
   791 
   792 ML {*
   793 store_met
   794  (prep_met thy "met_SP_Ztrans_inv" [] e_metID
   795  (["SignalProcessing", "Z_Transform", "inverse"], 
   796   [("#Given" ,["filterExpression X_eq"]),
   797    ("#Find"  ,["stepResponse n_eq"])
   798   ],
   799    {rew_ord'="tless_true", rls'= e_rls, calc = [], srls = srls, 
   800     prls = e_rls,
   801     crls = e_rls, nrls = e_rls},
   802       "Script InverseZTransform (X_eq::bool) =" ^ (*(1/z) instead of z ^^^ -1*)
   803       " (let X = Take X_eq;" ^
   804 (*([1], Frm), X z = 3 / (z - 1 / 4 + -1 / 8 * (1 / z))*)
   805       "      X' = Rewrite ruleZY False X;" ^ (*z * denominator*)
   806 (*([1], Res), ?X' z = 3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z)))*)
   807       "      (num_orig::real) = get_numerator (rhs X');"^
   808       "      X' = (Rewrite_Set norm_Rational False) X';" ^ (*simplify*)
   809 (*([2], Res), ?X' z = 24 / (-1 + -2 * z + 8 * z ^^^ 2)*)
   810       "      (X'_z::real) = lhs X';" ^ (**)
   811       "      (zzz::real) = argument_in X'_z;" ^ (**)
   812       "      (funterm::real) = rhs X';" ^ (*drop X' z = for equation solving*)
   813       "      (denom::real) = get_denominator funterm;" ^ (*get_denominator*)
   814       "      (num::real) = get_numerator funterm; " ^ (*get_numerator*)
   815       "      (equ::bool) = (denom = (0::real));" ^
   816       "      (L_L::bool list) = (SubProblem (PolyEq'," ^
   817       "         [abcFormula,degree_2,polynomial,univariate,equation],[no_met])" ^
   818       "         [BOOL equ, REAL zzz]);              " ^
   819 
   820 (*([3], Pbl), solve (-1 + -2 * z + 8 * z ^^^ 2 = 0, z)*)
   821 (*([3], Res), [z = 1 / 2, z = -1 / 4]*)
   822 
   823       "      (facs::real) = factors_from_solution L_L;" ^
   824       "      (eql::real) = Take (num_orig / facs);" ^                                 (*([4], Frm), 24 / ((z + -1 * (1 / 2)) * (z + -1 * (-1 / 4)))*)
   825 
   826       "      (eqr::real) = (Try (Rewrite_Set ansatz_rls False)) eql;"^           (*([4], Res), ?A / (z + -1 * (1 / 2)) + ?B / (z + -1 * (-1 / 4))*)
   827 
   828       "      (eq::bool) = Take (eql = eqr);"^ (*Maybe possible to use HOLogic.mk_eq ??*) (*([5], Frm), 24 / ((z + -1 * (1 / 2)) * (z + -1 * (-1 / 4))) = ?A / (z + -1 * (1 / 2)) + ?B / (z + -1 * (-1 / 4))*)
   829 
   830       "      eq = (Try (Rewrite_Set equival_trans False)) eq;"^                  (*([5], Res), 24 = ?A * (z + -1 * (-1 / 4)) + ?B * (z + -1 * (1 / 2))*)
   831 
   832       "      eq = drop_questionmarks eq;"^
   833       "      (z1::real) = (rhs (NTH 1 L_L));"^                                   (*prepare equliation for a - eq_a therfor subsitude z with solution 1 - z1*)
   834       "      (z2::real) = (rhs (NTH 2 L_L));"^
   835  
   836       "      (eq_a::bool) = Take eq;"^
   837       "      eq_a = (Substitute [zzz=z1]) eq;"^                          (*([6], Res), 24 = ?A * (1 / 2 + -1 * (-1 / 4)) + ?B * (1 / 2 + -1 * (1 / 2))*)
   838       "      eq_a = (Rewrite_Set norm_Rational False) eq_a;"^                    (*([7], Res), 24 = ?A * 3 / 4*)
   839       "      (sol_a::bool list) = (SubProblem (Isac'," ^
   840       "          [univariate,equation],[no_met])" ^
   841       "        [BOOL eq_a, REAL (A::real)]);"^
   842       "      (a::real) = (rhs(NTH 1 sol_a));"^
   843 
   844 
   845       "      (eq_b::bool) = Take eq;"^
   846       "      eq_b =  (Substitute [zzz=z2]) eq_b;"^
   847       "      eq_b = (Rewrite_Set norm_Rational False) eq_b;"^
   848       "      (sol_b::bool list) = (SubProblem (Isac'," ^
   849       "          [univariate,equation],[no_met])" ^
   850       "        [BOOL eq_b, REAL (B::real)]);"^
   851 
   852       "      (x::real) = Take num_orig"^(*only here that the last step of the ptree is visible in tracing output*)
   853       "  in X)" 
   854       ));
   855 *}
   856 
   857 
   858 subsection {*Check the Program*}
   859 
   860 subsubsection {*Check the formalization*}
   861 ML {*
   862 val fmz = ["filterExpression (X  = 3 / (z - 1/4 + -1/8 * (1/(z::real))))", 
   863   "stepResponse (x[n::real]::bool)"];
   864 val (dI,pI,mI) = ("Isac", ["inverse", "Z_Transform", "SignalProcessing"], 
   865   ["SignalProcessing","Z_Transform","inverse"]);
   866 
   867 val ([(1, [1], "#Given", Const ("Inverse_Z_Transform.filterExpression", _),
   868             [Const ("HOL.eq", _) $ _ $ _]),
   869            (2, [1], "#Find", Const ("Inverse_Z_Transform.stepResponse", _),
   870             [Free ("x", _) $ _])],
   871           _) = prep_ori fmz thy ((#ppc o get_pbt) pI);
   872 *}
   873 ML {*
   874 val Script sc = (#scr o get_met) ["SignalProcessing","Z_Transform","inverse"];
   875 atomty sc;
   876 *}
   877 
   878 subsubsection {*Stepwise check the program*}
   879 ML {*
   880 trace_rewrite := false;
   881 trace_script := false; print_depth 9;
   882 val fmz = ["filterExpression (X z = 3 / (z - 1/4 + -1/8 * (1/(z::real))))", 
   883   "stepResponse (x[n::real]::bool)"];
   884 val (dI,pI,mI) = ("Isac", ["inverse", "Z_Transform", "SignalProcessing"], 
   885   ["SignalProcessing","Z_Transform","inverse"]);
   886 val (p,_,f,nxt,_,pt)  = CalcTreeTEST [(fmz, (dI,pI,mI))];
   887 (*([], Frm), Problem (Isac, [inverse, Z_Transform, SignalProcessing])*)
   888 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "Add_Given";
   889 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "Add_Find";
   890 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "Specify_Theory";
   891 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "Specify_Problem";
   892 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "Specify_Method";
   893 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "nxt = Apply_Method";
   894 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!*)
   895 (*([1], Frm), X z = 3 / (z - 1 / 4 + -1 / 8 * (1 / z))*)
   896 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)))";
   897 (*([1], Res), ?X' z = 3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z)))*)
   898 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "nxt = SubProblem";
   899 (*([2], Res), ?X' z = 24 / (-1 + -2 * z + 8 * z ^^^ 2)*)
   900 *}
   901 text {* Instead of nxt = Subproblem above we had Empty_Tac; the search for the reason 
   902   considered the following points:
   903   # what shows show_pt pt; ...
   904     (([2], Res), ?X' z = 24 / (-1 + -2 * z + 8 * z ^^^ 2))] ..calculation ok,
   905     but no "next" step found: should be "nxt = Subproblem" ?!?
   906   # what shows trace_script := true; we read bottom up ...
   907     @@@ next   leaf 'SubProbfrom
   908      (PolyEq', [abcFormula, degree_2, polynomial, univariate, equation],
   909       no_meth)
   910      [BOOL equ, REAL z]' ---> STac 'SubProblem
   911      (PolyEq', [abcFormula, degree_2, polynomial, univariate, equation],
   912       no_meth)
   913      [BOOL (-1 + -2 * z + 8 * z ^^^ 2 = 0), REAL z]'
   914     ... and see the SubProblem with correct arguments from searching next step
   915     (program text !!!--->!!! STac (script tactic) with arguments evaluated.)
   916   # do we have the right Script ...difference in the argumentsdifference in the arguments
   917     val Script s = (#scr o get_met) ["SignalProcessing","Z_Transform","inverse"];
   918     writeln (term2str s);
   919     ... shows the right script.difference in the arguments
   920   # test --- why helpless here ? --- shows: replace no_meth by [no_meth] in Script
   921 *}
   922 
   923 ML {*
   924 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "nxt = Model_Problem";
   925 (*([3], Pbl), solve (-1 + -2 * z + 8 * z ^^^ 2 = 0, z)*)
   926 *}
   927 text {* Instead of nxt = Model_Problem above we had Empty_Tac; the search for the reason 
   928   considered the following points:difference in the arguments
   929   # comparison with subsection { *solve equation* }: there solving this equation works,
   930     so there must be some difference in the arguments of the Subproblem:
   931     RIGHT: we had [no_meth] instead of [no_met] ;-))
   932 *}
   933 ML {*
   934 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "nxt = Add_Given equality (-1 + -2 * z + 8 * z ^^^ 2 = 0)";
   935 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "nxt = Add_Given solveFor z";
   936 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "nxt = Add_Find solutions z_i";
   937 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "nxt = Specify_Theory Isac";
   938 *}
   939 
   940 text {* We had "nxt = Empty_Tac instead Specify_Theory; 
   941   the search for the reason considered the following points:
   942   # was there an error message ? NO --ok
   943   # has "nxt = Add_Find" been inserted in pt: get_obj g_pbl pt (fst p); YES --ok
   944   # what is the returned "formula": print_depth 999; f; print_depth 999; --
   945     {Find = [Correct "solutions z_i"], With = [], 
   946      Given = [Correct "equality (-1 + -2 * z + 8 * z ^^^ 2 = 0)", Correct "solveFor z"],
   947      Where = [False "matches (z = 0) (-1 + -2 * z + 8 * z ^^^ 2 = 0) |\n
   948                      matches (?b * z = 0) (-1 + -2 * z + 8 * z ^^^ 2 = 0) |\n
   949                      matches (?a + z = 0) (-1 + -2 * z + 8 * z ^^^ 2 = 0) |\n
   950                      matches (?a + ?b * z = 0) (-1 + -2 * z + 8 * z ^^^ 2 = 0)"],
   951      Relate = []}
   952      -- the only False is the reason: the Where (the precondition) is False for good reasons:
   953      the precondition seems to check for linear equations, not for the one we want to solve!
   954   Removed this error by correcting the Script
   955   from SubProblem (PolyEq', [linear,univariate,equation,test], [Test,solve_linear]
   956   to   SubProblem (PolyEq', [abcFormula,degree_2,polynomial,univariate,equation],
   957                    [PolyEq,solve_d2_polyeq_abc_equation]
   958   You find the appropriate type of equation at
   959     http://www.ist.tugraz.at/projects/isac/www/kbase/pbl/index_pbl.html
   960   and the respective method in Knowledge/PolyEq.thy at the respective store_pbt.
   961   Or you leave the selection of the appropriate type to isac as done in the final Script ;-))
   962 *}
   963 ML {*
   964 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "nxt = Specify_Problem [abcFormula,...";
   965 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "nxt = Specify_Method [PolyEq,solve_d2_polyeq_abc_equation";
   966 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "nxt = Apply_Method [PolyEq,solve_d2_polyeq_abc_equation";
   967 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "nxt = Rewrite_Set_Inst ([(bdv, z)], d2_polyeq_abcFormula_simplify";
   968 (*([3,1], Frm), -1 + -2 * z + 8 * z ^^^ 2 = 0)*)
   969 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   970 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   971 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   972 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   973 (*([3,4], Res), [z = 1 / 2, z = -1 / 4])*)
   974 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   975 (*([3], Res), [z = 1 / 2, z = -1 / 4]*)
   976 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   977 (*([4], Frm), 24 / ((z + -1 * (1 / 2)) * (z + -1 * (-1 / 4)))*)
   978 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   979 (*([4], Res), ?A / (z + -1 * (1 / 2)) + ?B / (z + -1 * (-1 / 4))*)
   980 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   981 (*([5], Frm), 24 / ((z + -1 * (1 / 2)) * (z + -1 * (-1 / 4))) = ?A / (z + -1 * (1 / 2)) + ?B / (z + -1 * (-1 / 4))*)
   982 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   983 (*([5], Res), 24 = ?A * (z + -1 * (-1 / 4)) + ?B * (z + -1 * (1 / 2))*)
   984 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   985 (*([6], Res), 24 = A * (1 / 2 + -1 * (-1 / 4)) + B * (1 / 2 + -1 * (1 / 2))*)
   986 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   987 (*([7], Res), 24 = A * 3 / 4*)
   988 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   989 (*([8], Pbl), solve (24 = 3 * A / 4, A)*)
   990 val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*Add_Given "equality (24 = 3 * A / 4)"*)
   991 val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*Add_Given "solveFor A"*)
   992 val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*Add_Find "solutions A_i"*)
   993 val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*Specify_Theory "Isac"*)val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*Specify_Problem ["normalize", "polynomial", 
   994                                          "univariate", "equation"]*)
   995 val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*Specify_Method ["PolyEq", "normalize_poly"]*)
   996 val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*Apply_Method ["PolyEq", "normalize_poly"]*)
   997 val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*Rewrite ("all_left", "PolyEq.all_left")*)
   998 val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*Rewrite_Set_Inst (["(bdv, A)"], "make_ratpoly_in")*)
   999 val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*Rewrite_Set "polyeq_simplify"*)
  1000 val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*Subproblem ("Isac", ["degree_1", "polynomial", 
  1001                                          "univariate", "equation"])*)
  1002 *}
  1003 
  1004 
  1005 ML {*
  1006 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1007 show_pt pt;
  1008 *}
  1009 
  1010 ML {*
  1011 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1012 show_pt pt;
  1013 *}
  1014 
  1015 ML {*
  1016 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1017 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1018 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1019 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1020 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1021 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1022 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1023 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1024 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1025 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1026 *}
  1027 
  1028 ML {*
  1029 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1030 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1031 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1032 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1033 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1034 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1035 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1036 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1037 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1038 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1039 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1040 *}
  1041 
  1042 ML {*
  1043 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1044 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1045 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1046 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1047 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1048 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1049 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1050 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1051 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1052 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1053 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1054 *}
  1055 
  1056 ML {*
  1057 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1058 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1059 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1060 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1061 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1062 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1063 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1064 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1065 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1066 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1067 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1068 *}
  1069 
  1070 ML {*
  1071 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1072 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1073 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1074 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1075 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1076 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1077 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1078 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1079 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1080 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1081 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1082 *}
  1083 
  1084 
  1085 ML {*
  1086 (*trace_script := true;*)
  1087 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1088 show_pt pt;
  1089 *}
  1090 
  1091 
  1092 section {*Write Tests for Crucial Details*}
  1093 text{*===================================*}
  1094 ML {*
  1095 *}
  1096 
  1097 section {*Integrate Program into Knowledge*}
  1098 ML {*
  1099 print_depth 999;
  1100 *}
  1101 
  1102 end
  1103