test/Tools/isac/ADDTESTS/course/SignalProcess/Build_Inverse_Z_Transform.thy
author wenzelm
Thu, 10 Jun 2021 12:23:57 +0200
changeset 60290 bb4e8b01b072
parent 60278 343efa173023
child 60297 73e7175a7d3f
permissions -rw-r--r--
clarified theory context: from current command instead of earlier state;
     1 (* Title:  Build_Inverse_Z_Transform
     2    Author: Jan Rocnik
     3    (c) copyright due to license terms.
     4 *)
     5 
     6 theory Build_Inverse_Z_Transform imports Isac.Inverse_Z_Transform
     7   
     8 begin
     9 
    10 text\<open>We stepwise build \ttfamily Inverse\_Z\_Transform.thy \normalfont as an 
    11   exercise. Because Subsection~\ref{sec:stepcheck} requires 
    12   \ttfamily Inverse\_Z\_Transform.thy \normalfont as a subtheory of \ttfamily 
    13   Isac.thy\normalfont, the setup has been changed from \ttfamily theory 
    14   Inverse\_Z\_Transform imports Isac \normalfont to the above one.
    15   \par \noindent
    16   \begin{center} 
    17   \textbf{Attention with the names of identifiers when going into internals!}
    18   \end{center}
    19   Here in this theory there are the internal names twice, for instance we have
    20   \ttfamily (Thm.derivation\_name @{thm rule1} = 
    21   "Build\_Inverse\_Z\_Transform.rule1") = true; \normalfont
    22   but actually in us will be \ttfamily Inverse\_Z\_Transform.rule1 \normalfont
    23 \<close>
    24 
    25 section \<open>Trials towards the Z-Transform\label{sec:trials}\<close>
    26 
    27 ML \<open>val thy = @{theory};\<close>
    28 
    29 subsection \<open>Notations and Terms\<close>
    30 text\<open>\noindent Try which notations we are able to use.\<close>
    31 ML \<open>
    32   @{term "1 < || z ||"};
    33   @{term "z / (z - 1)"};
    34   @{term "-u -n - 1"};
    35   @{term "-u [-n - 1]"}; (*[ ] denotes lists !!!*)
    36   @{term "z /(z - 1) = -u [-n - 1]"};
    37   @{term "1 < || z || ==> z / (z - 1) = -u [-n - 1]"};
    38   UnparseC.term @{term "1 < || z || ==> z / (z - 1) = -u [-n - 1]"};
    39 \<close>
    40 text\<open>\noindent Try which symbols we are able to use and how we generate them.\<close>
    41 ML \<open>
    42   (*alpha -->  "</alpha>" *)
    43   @{term "\<alpha> "};
    44   @{term "\<delta> "};
    45   @{term "\<phi> "};
    46   @{term "\<rho> "};
    47   UnparseC.term @{term "\<rho> "};
    48 \<close>
    49 
    50 subsection \<open>Rules\<close>
    51 (*axiomatization "z / (z - 1) = -u [-n - 1]"
    52   Illegal variable name: "z / (z - 1) = -u [-n - 1]" *)
    53 (*definition     "z / (z - 1) = -u [-n - 1]"
    54   Bad head of lhs: existing constant "op /"*)
    55 axiomatization where 
    56   rule1: "1 = \<delta>[n]" and
    57   rule2: "|| z || > 1 ==> z / (z - 1) = u [n]" and
    58   rule3: "|| z || < 1 ==> z / (z - 1) = -u [-n - 1]" and 
    59   rule4: "|| z || > || \<alpha> || ==> z / (z - \<alpha>) = \<alpha> \<up> n * u [n]" and
    60   rule5: "|| z || < || \<alpha> || ==> z / (z - \<alpha>) = -(\<alpha> \<up> n) * u [-n - 1]" and
    61   rule6: "|| z || > 1 ==> z/(z - 1) \<up> 2 = n * u [n]"
    62 
    63 text\<open>\noindent Check the rules for their correct notation. 
    64       (See the machine output.)\<close>
    65 ML \<open>
    66   @{thm rule1};
    67   @{thm rule2};
    68   @{thm rule3};
    69   @{thm rule4};
    70 \<close>
    71 
    72 subsection \<open>Apply Rules\<close>
    73 text\<open>\noindent We try to apply the rules to a given expression.\<close>
    74 
    75 ML \<open>
    76   val inverse_Z = Rule_Set.append_rules "inverse_Z" Atools_erls
    77     [ Rule.Thm  ("rule3",ThmC.numerals_to_Free @{thm rule3}),
    78       Rule.Thm  ("rule4",ThmC.numerals_to_Free @{thm rule4}),
    79       Rule.Thm  ("rule1",ThmC.numerals_to_Free @{thm rule1})   
    80     ];
    81 
    82 \<close> ML \<open>
    83   val t = TermC.str2term "z / (z - 1) + z / (z - \<alpha>) + 1";
    84 \<close> ML \<open>
    85   val SOME (t', asm) = Rewrite.rewrite_set_ thy true inverse_Z t;
    86   UnparseC.term t' = "z / (z - ?\<delta> [?n]) + z / (z - \<alpha>) + ?\<delta> [?n]";
    87   (*
    88    * Attention rule1 is applied before the expression is 
    89    * checked for rule4 which would be correct!!!
    90    *)
    91 \<close> ML \<open>
    92 \<close>
    93 
    94 ML \<open>val (thy, ro, er) = (@{theory}, tless_true, eval_rls);\<close>
    95 ML \<open>
    96   val SOME (t, asm1) = 
    97     Rewrite.rewrite_ thy ro er true (ThmC.numerals_to_Free @{thm rule3}) t;
    98   UnparseC.term t = "- ?u [- ?n - 1] + z / (z - \<alpha>) + 1";
    99   (*- real *)
   100   UnparseC.term t;
   101 
   102   val SOME (t, asm2) = 
   103     Rewrite.rewrite_ thy ro er true (ThmC.numerals_to_Free @{thm rule4}) t;
   104   UnparseC.term t = "- ?u [- ?n - 1] + \<alpha> \<up> ?n * ?u [?n] + 1";
   105   (*- real *)
   106   UnparseC.term t;
   107 
   108   val SOME (t, asm3) = 
   109     Rewrite.rewrite_ thy ro er true (ThmC.numerals_to_Free @{thm rule1}) t;
   110   UnparseC.term t = "- ?u [- ?n - 1] + \<alpha> \<up> ?n * ?u [?n] + ?\<delta> [?n]";
   111   (*- real *)
   112   UnparseC.term t;
   113 \<close>
   114 ML \<open>UnparseC.terms (asm1 @ asm2 @ asm3);\<close>
   115 
   116 section\<open>Prepare Steps for TP-based programming Language\label{sec:prepstep}\<close>
   117 text\<open>
   118       \par \noindent The following sections are challenging with the CTP-based 
   119       possibilities of building the program. The goal is realized in 
   120       Section~\ref{spec-meth} and Section~\ref{prog-steps}.
   121       \par The reader is advised to jump between the subsequent subsections and 
   122       the respective steps in Section~\ref{prog-steps}. By comparing 
   123       Section~\ref{sec:calc:ztrans} the calculation can be comprehended step 
   124       by step.
   125 \<close>
   126 
   127 subsection \<open>Prepare Expression\label{prep-expr}\<close>
   128 text\<open>\noindent We try two different notations and check which of them 
   129        Isabelle can handle best.\<close>
   130 ML \<open>
   131   val ctxt = Proof_Context.init_global @{theory};
   132 (*val ctxt = ContextC.declare_constraints' [@{term "z::real"}] ctxt;*)
   133 
   134   val SOME fun1 = 
   135     TermC.parseNEW ctxt "X z = 3 / (z - 1/4 + -1/8 * z \<up> -1)"; UnparseC.term fun1;
   136   val SOME fun1' = 
   137     TermC.parseNEW ctxt "X z = 3 / (z - 1/4 + -1/8 * (1/z))"; UnparseC.term fun1';
   138 \<close>
   139 
   140 subsubsection \<open>Prepare Numerator and Denominator\<close>
   141 text\<open>\noindent The partial fraction decomposition is only possible if we
   142        get the bound variable out of the numerator. Therefor we divide
   143        the expression by $z$. Follow up the Calculation at 
   144        Section~\ref{sec:calc:ztrans} line number 02.\<close>
   145 
   146 axiomatization where
   147   ruleZY: "(X z = a / b) = (X' z = a / (z * b))"
   148 
   149 ML \<open>
   150   val (thy, ro, er) = (@{theory}, tless_true, eval_rls);
   151   val SOME (fun2, asm1) = 
   152     Rewrite.rewrite_ thy ro er true  @{thm ruleZY} fun1; UnparseC.term fun2;
   153   val SOME (fun2', asm1) = 
   154     Rewrite.rewrite_ thy ro er true  @{thm ruleZY} fun1'; UnparseC.term fun2';
   155 
   156   val SOME (fun3,_) = 
   157     Rewrite.rewrite_set_ @{theory} false norm_Rational fun2;
   158   UnparseC.term fun3;
   159   (*
   160    * Fails on x \<up> (-1)
   161    * We solve this problem by using 1/x as a workaround.
   162    *)
   163   val SOME (fun3',_) = 
   164     Rewrite.rewrite_set_ @{theory} false norm_Rational fun2';
   165   UnparseC.term fun3';
   166   (*
   167    * OK - workaround!
   168    *)
   169 \<close>
   170 
   171 subsubsection \<open>Get the Argument of the Expression X'\<close>
   172 text\<open>\noindent We use \texttt{grep} for finding possibilities how we can
   173        extract the bound variable in the expression. \ttfamily Prog_Expr.thy, 
   174        Tools.thy \normalfont contain general utilities: \ttfamily 
   175        eval\_argument\_in, eval\_rhs, eval\_lhs,\ldots \normalfont
   176        \ttfamily grep -r "fun eva\_" * \normalfont shows all functions 
   177        witch can be used in a script. Lookup this files how to build 
   178        and handle such functions.
   179        \par The next section shows how to introduce such a function.
   180 \<close>
   181 
   182 subsubsection\<open>Decompose the Given Term Into lhs and rhs\<close>
   183 ML \<open>
   184   val (_, expr) = HOLogic.dest_eq fun3'; UnparseC.term expr;
   185   val (_, denom) = 
   186     HOLogic.dest_bin "Rings.divide_class.divide" (type_of expr) expr;
   187   UnparseC.term denom = "-1 + -2 * z + 8 * z \<up> 2";
   188 \<close>
   189 
   190 text\<open>\noindent We have rhs\footnote{Note: lhs means \em Left Hand Side
   191       \normalfont and rhs means \em Right Hand Side \normalfont and indicates
   192       the left or the right part of an equation.} in the Program language, but
   193       we need a function which gets the denominator of a fraction.\<close>
   194 
   195 subsubsection\<open>Get the Denominator and Numerator out of a Fraction\<close>
   196 text\<open>\noindent The self written functions in e.g. \texttt{get\_denominator}
   197        should become a constant for the Isabelle parser:\<close>
   198 
   199 consts
   200   get_denominator :: "real => real"
   201   get_numerator :: "real => real"
   202 
   203 text \<open>\noindent With the above definition we run into problems when we parse
   204         the Program \texttt{InverseZTransform}. This leads to \em ambiguous
   205         parse trees. \normalfont We avoid this by moving the definition
   206         to \ttfamily Rational.thy \normalfont and re-building {\sisac}.
   207         \par \noindent ATTENTION: From now on \ttfamily 
   208         Build\_Inverse\_Z\_Transform \normalfont mimics a build from scratch;
   209         it only works due to re-building {\sisac} several times (indicated 
   210         explicitly).
   211 \<close>
   212 
   213 ML \<open>
   214 (*
   215  *("get_denominator",
   216  *  ("Rational.get_denominator", eval_get_denominator ""))
   217  *)
   218 fun eval_get_denominator (thmid:string) _ 
   219 		      (t as Const ("Rational.get_denominator", _) $
   220               (Const ("Rings.divide_class.divide", _) $num 
   221                 $denom)) thy = 
   222         SOME (TermC.mk_thmid thmid (UnparseC.term_in_thy thy denom) "", 
   223 	        HOLogic.Trueprop $ (TermC.mk_equality (t, denom)))
   224   | eval_get_denominator _ _ _ _ = NONE; 
   225 \<close>
   226 text \<open>\noindent For the tests of \ttfamily eval\_get\_denominator \normalfont
   227         see \ttfamily test/Knowledge/rational.sml\normalfont\<close>
   228 
   229 text \<open>\noindent \ttfamily get\_numerator \normalfont should also become a
   230         constant for the Isabelle parser, follow up the \texttt{const}
   231         declaration above.\<close>
   232 
   233 ML \<open>
   234 (*
   235  *("get_numerator",
   236  *  ("Rational.get_numerator", eval_get_numerator ""))
   237  *)
   238 fun eval_get_numerator (thmid:string) _ 
   239 		      (t as Const ("Rational.get_numerator", _) $
   240               (Const ("Rings.divide_class.divide", _) $num
   241                 $denom )) thy = 
   242         SOME (TermC.mk_thmid thmid (UnparseC.term_in_thy thy num) "", 
   243 	        HOLogic.Trueprop $ (TermC.mk_equality (t, num)))
   244   | eval_get_numerator _ _ _ _ = NONE; 
   245 \<close>
   246 
   247 text \<open>\noindent We discovered several problems by implementing the 
   248        \ttfamily get\_numerator \normalfont function. Remember when 
   249        putting new functions to {\sisac}, put them in a thy file and rebuild 
   250        {\sisac}, also put them in the ruleset for the script!\<close>
   251 
   252 subsection \<open>Solve Equation\label{sec:solveq}\<close>
   253 text \<open>\noindent We have to find the zeros of the term, therefor we use our
   254        \ttfamily get\_denominator \normalfont function from the step before
   255        and try to solve the second order equation. (Follow up the Calculation
   256        in Section~\ref{sec:calc:ztrans} Subproblem 2) Note: This type of
   257        equation is too general for the present program.
   258        \par We know that this equation can be categorized as \em univariate
   259        equation \normalfont and solved with the functions {\sisac} provides
   260        for this equation type. Later on {\sisac} should determine the type
   261        of the given equation self.\<close>
   262 ML \<open>
   263   val denominator = TermC.parseNEW ctxt "z \<up> 2 - 1/4*z - 1/8 = 0";
   264   val fmz = ["equality (z \<up> 2 - 1/4*z - 1/8 = (0::real))",
   265              "solveFor z",
   266              "solutions L"];
   267   val (dI',pI',mI') =("Isac_Knowledge", ["univariate", "equation"], ["no_met"]);
   268 \<close>
   269 text \<open>\noindent Check if the given equation matches the specification of this
   270         equation type.\<close>
   271 ML \<open>
   272   M_Match.match_pbl fmz (Problem.from_store ["univariate", "equation"]);
   273 \<close>
   274 
   275 text\<open>\noindent We switch up to the {\sisac} Context and try to solve the 
   276        equation with a more specific type definition.\<close>
   277 
   278 ML \<open>
   279   Context.theory_name thy = "Isac_Knowledge";
   280   val denominator = TermC.parseNEW ctxt "-1 + -2 * z + 8 * z \<up> 2 = 0";
   281   val fmz =                                             (*specification*)
   282     ["equality (-1 + -2 * z + 8 * z \<up> 2 = (0::real))",(*equality*)
   283      "solveFor z",                                      (*bound variable*)
   284      "solutions L"];                                    (*identifier for
   285                                                           solution*)
   286   val (dI',pI',mI') =
   287     ("Isac_Knowledge", 
   288       ["abcFormula", "degree_2", "polynomial", "univariate", "equation"],
   289       ["no_met"]);
   290 \<close>
   291 
   292 text \<open>\noindent Check if the (other) given equation matches the 
   293         specification of this equation type.\<close>
   294         
   295 ML \<open>
   296   M_Match.match_pbl fmz (Problem.from_store
   297     ["abcFormula", "degree_2", "polynomial", "univariate", "equation"]);
   298 \<close>
   299 
   300 text \<open>\noindent We stepwise solve the equation. This is done by the
   301         use of a so called calc tree seen downwards.\<close>
   302 
   303 ML \<open>
   304   val (p,_,f,nxt,_,pt) = CalcTreeTEST [(fmz, (dI',pI',mI'))];
   305   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   306   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   307   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   308   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   309   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   310   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   311   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   312   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   313   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   314   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   315   val (p,_,f,nxt,_,pt) = me nxt p [] pt;         
   316   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   317   (*
   318    * nxt =..,Check_elementwise "Assumptions") 
   319    *)
   320   val (p,_,f,nxt,_,pt) = me nxt p [] pt;         
   321   val (p,_,f,nxt,_,pt) = me nxt p [] pt; f2str f;
   322   (*
   323    * [z = 1 / 2, z = -1 / 4]
   324    *)
   325   Test_Tool.show_pt pt; 
   326   val SOME f = parseNEW ctxt "[z=1/2, z=-1/4]";
   327 \<close>
   328 
   329 subsection \<open>Partial Fraction Decomposition\label{sec:pbz}\<close>
   330 text\<open>\noindent We go on with the decomposition of our expression. Follow up the
   331        Calculation in Section~\ref{sec:calc:ztrans} Step~3 and later on
   332        Subproblem~1.\<close>
   333 subsubsection \<open>Solutions of the Equation\<close>
   334 text\<open>\noindent We get the solutions of the before solved equation in a list.\<close>
   335 
   336 ML \<open>
   337   val SOME solutions = parseNEW ctxt "[z=1/2, z=-1/4]";
   338   UnparseC.term solutions;
   339   atomty solutions;
   340 \<close>
   341 
   342 subsubsection \<open>Get Solutions out of a List\<close>
   343 text \<open>\noindent In {\sisac}'s TP-based programming language: 
   344 \begin{verbatim}
   345   let $ $ s_1 = NTH 1 $ solutions; $ s_2 = NTH 2... $
   346 \end{verbatim}
   347        can be useful.
   348 \<close>
   349 
   350 ML \<open>
   351   val Const ("List.list.Cons", _) $ s_1 $ (Const ("List.list.Cons", _)
   352         $ s_2 $ Const ("List.list.Nil", _)) = solutions;
   353   UnparseC.term s_1;
   354   UnparseC.term s_2;
   355 \<close>
   356 
   357 text\<open>\noindent The ansatz for the \em Partial Fraction Decomposition \normalfont
   358       requires to get the denominators of the partial fractions out of the 
   359       Solutions as:
   360       \begin{itemize}
   361         \item $Denominator_{1}=z-Zeropoint_{1}$
   362         \item $Denominator_{2}=z-Zeropoint_{2}$
   363         \item \ldots
   364       \end{itemize}
   365 \<close>
   366       
   367 ML \<open>
   368   val xx = HOLogic.dest_eq s_1;
   369   val s_1' = HOLogic.mk_binop "Groups.minus_class.minus" xx;
   370   val xx = HOLogic.dest_eq s_2;
   371   val s_2' = HOLogic.mk_binop "Groups.minus_class.minus" xx;
   372   UnparseC.term s_1';
   373   UnparseC.term s_2';
   374 \<close>
   375 
   376 text \<open>\noindent For the programming language a function collecting all the 
   377         above manipulations is helpful.\<close>
   378 
   379 ML \<open>
   380   fun fac_from_sol s =
   381     let val (lhs, rhs) = HOLogic.dest_eq s
   382     in HOLogic.mk_binop "Groups.minus_class.minus" (lhs, rhs) end;
   383 \<close>
   384 
   385 ML \<open>
   386   fun mk_prod prod [] =
   387         if prod = TermC.empty
   388         then error "mk_prod called with []" 
   389         else prod
   390     | mk_prod prod (t :: []) =
   391         if prod = TermC.empty
   392         then t
   393         else HOLogic.mk_binop "Groups.times_class.times" (prod, t)
   394     | mk_prod prod (t1 :: t2 :: ts) =
   395           if prod = TermC.empty 
   396           then 
   397              let val p = 
   398                HOLogic.mk_binop "Groups.times_class.times" (t1, t2)
   399              in mk_prod p ts end 
   400           else 
   401              let val p =
   402                HOLogic.mk_binop "Groups.times_class.times" (prod, t1)
   403              in mk_prod p (t2 :: ts) end 
   404 \<close>
   405 (* ML {* 
   406 probably keep these test in test/Tools/isac/...
   407 (*mk_prod TermC.empty [];*)
   408 
   409 val prod = mk_prod TermC.empty [str2term "x + 123"]; 
   410 UnparseC.term prod = "x + 123";
   411 
   412 val sol = str2term "[z = 1 / 2, z = -1 / 4]";
   413 val sols = HOLogic.dest_list sol;
   414 val facs = map fac_from_sol sols;
   415 val prod = mk_prod TermC.empty facs; 
   416 UnparseC.term prod = "(z + -1 * (1 / 2)) * (z + -1 * (-1 / 4))";
   417 
   418 val prod = 
   419   mk_prod TermC.empty [str2term "x + 1", str2term "x + 2", str2term "x + 3"]; 
   420 UnparseC.term prod = "(x + 1) * (x + 2) * (x + 3)";
   421 *} *)
   422 ML \<open>
   423   fun factors_from_solution sol = 
   424     let val ts = HOLogic.dest_list sol
   425     in mk_prod TermC.empty (map fac_from_sol ts) end;
   426 \<close>
   427 (* ML {*
   428 val sol = str2term "[z = 1 / 2, z = -1 / 4]";
   429 val fs = factors_from_solution sol;
   430 UnparseC.term fs = "(z + -1 * (1 / 2)) * (z + -1 * (-1 / 4))"
   431 *} *)
   432 text \<open>\noindent This function needs to be packed such that it can be evaluated
   433         by the Lucas-Interpreter. Therefor we moved the function to the
   434         corresponding \ttfamily Equation.thy \normalfont in our case
   435         \ttfamily PartialFractions.thy \normalfont. The necessary steps
   436         are quit the same as we have done with \ttfamily get\_denominator 
   437         \normalfont before.\<close>
   438 ML \<open>
   439   (*("factors_from_solution",
   440     ("Partial_Fractions.factors_from_solution",
   441       eval_factors_from_solution ""))*)
   442       
   443   fun eval_factors_from_solution (thmid:string) _
   444       (t as Const ("Partial_Fractions.factors_from_solution", _) $ sol) thy =
   445         let val prod = factors_from_solution sol
   446         in SOME (TermC.mk_thmid thmid (UnparseC.term_in_thy thy prod) "",
   447              HOLogic.Trueprop $ (TermC.mk_equality (t, prod)))
   448         end
   449     | eval_factors_from_solution _ _ _ _ = NONE;
   450 \<close>
   451 
   452 text \<open>\noindent The tracing output of the calc tree after applying this
   453        function was:
   454 \begin{verbatim}
   455   24 / factors_from_solution [z = 1/ 2, z = -1 / 4])]
   456 \end{verbatim}
   457        and the next step:
   458 \begin{verbatim}
   459   val nxt = ("Empty_Tac", ...): tac'_)
   460 \end{verbatim}
   461        These observations indicate, that the Lucas-Interpreter (LIP) 
   462        does not know how to evaluate \ttfamily factors\_from\_solution
   463        \normalfont, so we knew that there is something wrong or missing.
   464 \<close>
   465        
   466 text\<open>\noindent First we isolate the difficulty in the program as follows:
   467 \begin{verbatim}      
   468   " (L_L::bool list) = (SubProblem (PolyEqX,      " ^
   469   "   [abcFormula, degree_2, polynomial,          " ^
   470   "    univariate,equation],                      " ^
   471   "   [no_met])                                   " ^
   472   "   [BOOL equ, REAL zzz]);                      " ^
   473   " (facs::real) = factors_from_solution L_L;     " ^
   474   " (foo::real) = Take facs                       " ^
   475 \end{verbatim}
   476 
   477       \par \noindent And see the tracing output:
   478       
   479 \begin{verbatim}
   480   [(([], Frm), Problem (Isac, [inverse, 
   481                                Z_Transform,
   482                                 SignalProcessing])),
   483    (([1], Frm), X z = 3 / (z - 1 / 4 + -1 / 8 * (1 / z))),
   484    (([1], Res), ?X' z = 3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z)))),
   485    (([2], Res), ?X' z = 24 / (-1 + -2 * z + 8 * z \<up> 2)),
   486    (([3], Pbl), solve (-1 + -2 * z + 8 * z \<up> 2 = 0, z)),
   487    (([3,1], Frm), -1 + -2 * z + 8 * z \<up> 2 = 0),
   488    (([3,1], Res), z = (- -2 + sqrt (-2 \<up> 2 - 4 * 8 * -1)) / (2 * 8)|
   489                   z = (- -2 - sqrt (-2 \<up> 2 - 4 * 8 * -1)) / (2 * 8)),
   490    (([3,2], Res), z = 1 / 2 | z = -1 / 4),
   491    (([3,3], Res), [ z = 1 / 2, z = -1 / 4]),
   492    (([3,4], Res), [ z = 1 / 2, z = -1 / 4]),
   493    (([3], Res), [ z = 1 / 2, z = -1 / 4]),
   494    (([4], Frm), factors_from_solution [z = 1 / 2, z = -1 / 4])]
   495 \end{verbatim}      
   496       
   497       \par \noindent In particular that:
   498       
   499 \begin{verbatim}
   500   (([3], Pbl), solve (-1 + -2 * z + 8 * z \<up> 2 = 0, z)),
   501 \end{verbatim}
   502       \par \noindent Shows the equation which has been created in
   503       the program by: 
   504 \begin{verbatim}
   505   "(denom::real) = get_denominator funterm;      " ^ 
   506     (* get_denominator *)
   507   "(equ::bool) = (denom = (0::real));            " ^
   508 \end{verbatim}
   509         
   510       \noindent \ttfamily get\_denominator \normalfont has been evaluated successfully,
   511       but not\\ \ttfamily factors\_from\_solution.\normalfont
   512       So we stepwise compare with an analogous case, \ttfamily get\_denominator
   513       \normalfont successfully done above: We know that LIP evaluates
   514       expressions in the program by use of the \emph{srls}, so we try to get
   515       the original \emph{srls}.\\
   516 
   517 \begin{verbatim}
   518   val {srls,...} = MethodC.from_store ["SignalProcessing",
   519                             "Z_Transform",
   520                             "Inverse"];
   521 \end{verbatim}
   522           
   523       \par \noindent Create 2 good example terms:
   524 
   525 \begin{verbatim}
   526 val SOME t1 =
   527   parseNEW ctxt "get_denominator ((111::real) / 222)";
   528 val SOME t2 =
   529   parseNEW ctxt "factors_from_solution [(z::real)=1/2, z=-1/4]";
   530 \end{verbatim}
   531 
   532       \par \noindent Rewrite the terms using srls:\\
   533       \ttfamily \par \noindent rewrite\_set\_ thy true srls t1;\\
   534         rewrite\_set\_ thy true srls t2;\\
   535       \par \noindent \normalfont Now we see a difference: \texttt{t1} gives
   536       \texttt{SOME} but \texttt{t2} gives \texttt{NONE}. We look at the 
   537       \emph{srls}:
   538 \begin{verbatim}
   539   val srls = 
   540     Rule_Set.Repeat{id = "srls_InverseZTransform",
   541         rules = [Eval("Rational.get_numerator",
   542                    eval_get_numerator "Rational.get_numerator"),
   543                  Eval("Partial_Fractions.factors_from_solution",
   544                    eval_factors_from_solution 
   545                      "Partial_Fractions.factors_from_solution")]}
   546 \end{verbatim}                
   547       \par \noindent Here everthing is perfect. So the error can
   548       only be in the SML code of \ttfamily eval\_factors\_from\_solution.
   549       \normalfont We try to check the code with an existing test; since the 
   550       \emph{code} is in 
   551       \begin{center}\ttfamily src/Tools/isac/Knowledge/Partial\_Fractions.thy
   552       \normalfont\end{center}
   553       the \emph{test} should be in
   554       \begin{center}\ttfamily test/Tools/isac/Knowledge/partial\_fractions.sml
   555       \normalfont\end{center}
   556       \par \noindent After updating the function \ttfamily
   557       factors\_from\_solution \normalfont to a new version and putting a
   558       test-case to \ttfamily Partial\_Fractions.sml \normalfont we tried again
   559       to evaluate the term with the same result.
   560       \par We opened the test \ttfamily Test\_Isac.thy \normalfont and saw that
   561       everything is working fine. Also we checked that the test \ttfamily 
   562       partial\_fractions.sml \normalfont is used in \ttfamily Test\_Isac.thy 
   563       \normalfont
   564       \begin{center}use \ttfamily "Knowledge/partial\_fractions.sml"
   565       \normalfont \end{center}
   566       and \ttfamily Partial\_Fractions.thy \normalfont is part is part of
   567       {\sisac} by evaluating
   568 
   569 \begin{verbatim}
   570   val thy = @{theory "Inverse_Z_Transform"};
   571 \end{verbatim}
   572 
   573       After rebuilding {\sisac} again it worked.
   574 \<close>
   575 
   576 subsubsection \<open>Build Expression\<close>
   577 text \<open>\noindent In {\sisac}'s TP-based programming language we can build
   578        expressions by:\\
   579        \ttfamily let s\_1 = Take numerator / (s\_1 * s\_2) \normalfont\<close>
   580        
   581 ML \<open>
   582   (*
   583    * The main denominator is the multiplication of the denominators of
   584    * all partial fractions.
   585    *)
   586    
   587   val denominator' = HOLogic.mk_binop 
   588     "Groups.times_class.times" (s_1', s_2') ;
   589   val SOME numerator = parseNEW ctxt "3::real";
   590 
   591   val expr' = HOLogic.mk_binop
   592     "Rings.divide_class.divide" (numerator, denominator');
   593   UnparseC.term expr';
   594 \<close>
   595 
   596 subsubsection \<open>Apply the Partial Fraction Decomposion Ansatz\<close>
   597 
   598 text\<open>\noindent We use the Ansatz of the Partial Fraction Decomposition for our
   599       expression 2nd order. Follow up the calculation in 
   600       Section~\ref{sec:calc:ztrans} Step~03.\<close>
   601 
   602 ML \<open>Context.theory_name thy = "Isac_Knowledge"\<close>
   603 
   604 text\<open>\noindent We define two axiomatization, the first one is the main ansatz,
   605       the next one is just an equivalent transformation of the resulting
   606       equation. Both axiomatizations were moved to \ttfamily
   607       Partial\_Fractions.thy \normalfont and got their own rulesets. In later
   608       programs it is possible to use the rulesets and the machine will find
   609       the correct ansatz and equivalent transformation itself.\<close>
   610 
   611 axiomatization where
   612   ansatz_2nd_order: "n / (a*b) = A/a + B/b" and
   613   equival_trans_2nd_order: "(n/(a*b) = A/a + B/b) = (n = A*b + B*a)"
   614 
   615 text\<open>\noindent We use our \ttfamily ansatz\_2nd\_order \normalfont to rewrite
   616        our expression and get an equation with our expression on the left
   617        and the partial fractions of it on the right hand side.\<close>
   618   
   619 ML \<open>
   620   val SOME (t1,_) = 
   621     rewrite_ @{theory} e_rew_ord Rule_Set.empty false 
   622              @{thm ansatz_2nd_order} expr';
   623   UnparseC.term t1; atomty t1;
   624   val eq1 = HOLogic.mk_eq (expr', t1);
   625   UnparseC.term eq1;
   626 \<close>
   627 
   628 text\<open>\noindent Eliminate the denominators by multiplying the left and the
   629       right hand side of the equation with the main denominator. This is an
   630       simple equivalent transformation. Later on we use an own ruleset
   631       defined in \ttfamily Partial\_Fractions.thy \normalfont for doing this.
   632       Follow up the calculation in Section~\ref{sec:calc:ztrans} Step~04.\<close>
   633 
   634 ML \<open>
   635   val SOME (eq2,_) = 
   636     rewrite_ @{theory} e_rew_ord Rule_Set.empty false 
   637              @{thm equival_trans_2nd_order} eq1;
   638   UnparseC.term eq2;
   639 \<close>
   640 
   641 text\<open>\noindent We use the existing ruleset \ttfamily norm\_Rational \normalfont 
   642      for simplifications on expressions.\<close>
   643 
   644 ML \<open>
   645   val SOME (eq3,_) = rewrite_set_ @{theory} false norm_Rational eq2;
   646   UnparseC.term eq3;
   647   (*
   648    * ?A ?B not simplified
   649    *)
   650 \<close>
   651 
   652 text\<open>\noindent In Example~\ref{eg:gap} of my thesis I'm describing a problem about
   653       simplifications. The problem that we would like to have only a specific degree
   654       of simplification occurs right here, in the next step.\<close>
   655 
   656 ML \<open>
   657   Rewrite.trace_on := false;
   658   val SOME fract1 =
   659     parseNEW ctxt "(z - 1/2)*(z - -1/4) * (A/(z - 1/2) + B/(z - -1/4))";
   660   (*
   661    * A B !
   662    *)
   663   val SOME (fract2,_) = 
   664     rewrite_set_ @{theory} false norm_Rational fract1;
   665   UnparseC.term fract2 = "(A + -2 * B + 4 * A * z + 4 * B * z) / 4";
   666   (*
   667    * UnparseC.term fract2 = "A * (1 / 4 + z) + B * (-1 / 2 + z)" 
   668    * would be more traditional...
   669    *)
   670 \<close>
   671 
   672 text\<open>\noindent We walk around this problem by generating our new equation first.\<close>
   673 
   674 ML \<open>
   675   val (numerator, denominator) = HOLogic.dest_eq eq3;
   676   val eq3' = HOLogic.mk_eq (numerator, fract1);
   677   (*
   678    * A B !
   679    *)
   680   UnparseC.term eq3';
   681   (*
   682    * MANDATORY: simplify (and remove denominator) otherwise 3 = 0
   683    *)
   684   val SOME (eq3'' ,_) = 
   685     rewrite_set_ @{theory} false norm_Rational eq3';
   686   UnparseC.term eq3'';
   687 \<close>
   688 
   689 text\<open>\noindent Still working at {\sisac}\ldots\<close>
   690 
   691 ML \<open>Context.theory_name thy = "Isac_Knowledge"\<close>
   692 
   693 subsubsection \<open>Build a Rule-Set for the Ansatz\<close>
   694 text \<open>\noindent The \emph{ansatz} rules violate the principle that each
   695        variable on the right-hand-side must also occur on the
   696        left-hand-side of the rule: A, B, etc. don't do that. Thus the
   697        rewriter marks these variables with question marks: ?A, ?B, etc.
   698        These question marks can be dropped by \ttfamily fun
   699        drop\_questionmarks\normalfont.\<close>
   700        
   701 ML \<open>
   702   val ansatz_rls = prep_rls @{theory} (
   703     Rule_Set.Repeat {id = "ansatz_rls", preconds = [], rew_ord = ("dummy_ord",dummy_ord),
   704       erls = Rule_Set.empty, srls = Rule_Set.Empty, calc = [], errpatts = [],
   705       rules = [
   706         Thm ("ansatz_2nd_order",ThmC.numerals_to_Free @{thm ansatz_2nd_order}),
   707         Thm ("equival_trans_2nd_order",ThmC.numerals_to_Free @{thm equival_trans_2nd_order})
   708               ], 
   709       scr = Empty_Prog});
   710 \<close>
   711 
   712 text\<open>\noindent We apply the ruleset\ldots\<close>
   713 
   714 ML \<open>
   715   val SOME (ttttt,_) = 
   716     rewrite_set_ @{theory} false ansatz_rls expr';
   717 \<close>
   718 
   719 text\<open>\noindent And check the output\ldots\<close>
   720 
   721 ML \<open>
   722   UnparseC.term expr' = "3 / ((z - 1 / 2) * (z - -1 / 4))";
   723   UnparseC.term ttttt = "?A / (z - 1 / 2) + ?B / (z - -1 / 4)";
   724 \<close>
   725 
   726 subsubsection \<open>Get the First Coefficient\<close>
   727 
   728 text\<open>\noindent Now it's up to get the two coefficients A and B, which will be
   729       the numerators of our partial fractions. Continue following up the 
   730       Calculation in Section~\ref{sec:calc:ztrans} Subproblem~1.\<close>
   731       
   732 text\<open>\noindent To get the first coefficient we substitute $z$ with the first
   733       zero-point we determined in Section~\ref{sec:solveq}.\<close>
   734 
   735 ML \<open>
   736   val SOME (eq4_1,_) =
   737     rewrite_terms_ @{theory} e_rew_ord Rule_Set.empty [s_1] eq3'';
   738   UnparseC.term eq4_1;
   739   val SOME (eq4_2,_) =
   740     rewrite_set_ @{theory} false norm_Rational eq4_1;
   741   UnparseC.term eq4_2;
   742 
   743   val fmz = ["equality (3=3*A/(4::real))", "solveFor A", "solutions L"];
   744   val (dI',pI',mI') =("Isac_Knowledge", ["univariate", "equation"], ["no_met"]);
   745   (*
   746    * Solve the simple linear equation for A:
   747    * Return eq, not list of eq's
   748    *)
   749   val (p,_,fa,nxt,_,pt) = CalcTreeTEST [(fmz, (dI',pI',mI'))];
   750   val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   751     (*Add_Given "equality (3=3*A/4)"*)
   752   val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   753     (*Add_Given "solveFor A"*)
   754   val (p,_,fa,nxt,_,pt) = me nxt p [] pt; 
   755     (*Add_Find "solutions L"*)
   756   val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   757     (*Specify_Theory "Isac_Knowledge"*)
   758   val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   759     (*Specify_Problem ["normalise", "polynomial",
   760                        "univariate", "equation"])*)
   761   val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   762     (* Specify_Method["PolyEq", "normalise_poly"]*)
   763   val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   764     (*Apply_Method["PolyEq", "normalise_poly"]*)
   765   val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   766     (*Rewrite ("all_left", "PolyEq.all_left")*)
   767   val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   768     (*Rewrite_Set_Inst(["(''bdv'',A)"],"make_ratpoly_in")*)
   769   val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   770     (*Rewrite_Set "polyeq_simplify"*)
   771   val (p,_,fa,nxt,_,pt) = me nxt p [] pt; 
   772   val (p,_,fa,nxt,_,pt) = me nxt p [] pt; 
   773   val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   774     (*Add_Given "equality (3 + -3 / 4 * A =0)"*)
   775   val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   776     (*Add_Given "solveFor A"*)
   777   val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   778     (*Add_Find "solutions A_i"*)
   779   val (p,_,fa,nxt,_,pt) = me nxt p [] pt; 
   780   val (p,_,fa,nxt,_,pt) = me nxt p [] pt; 
   781   val (p,_,fa,nxt,_,pt) = me nxt p [] pt; 
   782   val (p,_,fa,nxt,_,pt) = me nxt p [] pt; 
   783     (*Apply_Method ["PolyEq", "solve_d1_polyeq_equation"]*)
   784   val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   785     (*Rewrite_Set_Inst(["(''bdv'',A)"],"d1_polyeq_simplify")*)
   786   val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   787     (*Rewrite_Set "polyeq_simplify"*)
   788   val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   789     (*Rewrite_Set "norm_Rational_parenthesized"*)
   790   val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   791     (*Or_to_List*)
   792   val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   793     (*Check_elementwise "Assumptions"*)
   794   val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   795     (*Check_Postcond ["degree_1", "polynomial",
   796                       "univariate", "equation"]*)
   797   val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   798     (*Check_Postcond ["normalise", "polynomial",
   799                       "univariate", "equation"]*)
   800   val (p,_,fa,nxt,_,pt) = me nxt p [] pt;
   801     (*End_Proof'*)
   802   f2str fa;
   803 \<close>
   804 
   805 subsubsection \<open>Get Second Coefficient\<close>
   806 
   807 text\<open>\noindent With the use of \texttt{thy} we check which theories the 
   808       interpreter knows.\<close>
   809 
   810 ML \<open>thy\<close>
   811 
   812 text\<open>\noindent To get the second coefficient we substitute $z$ with the second
   813       zero-point we determined in Section~\ref{sec:solveq}.\<close>
   814 
   815 ML \<open>
   816   val SOME (eq4b_1,_) =
   817     rewrite_terms_ @{theory} e_rew_ord Rule_Set.empty [s_2] eq3'';
   818   UnparseC.term eq4b_1;
   819   val SOME (eq4b_2,_) =
   820     rewrite_set_ @{theory} false norm_Rational eq4b_1;
   821   UnparseC.term eq4b_2;
   822 
   823   val fmz = ["equality (3= -3*B/(4::real))", "solveFor B", "solutions L"];
   824   val (dI',pI',mI') =("Isac_Knowledge", ["univariate", "equation"], ["no_met"]);
   825   val (p,_,fb,nxt,_,pt) = CalcTreeTEST [(fmz, (dI',pI',mI'))];
   826   val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   827   val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   828   val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   829   val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   830   val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   831   val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   832   val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   833   val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   834   val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   835   val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   836   val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   837   val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   838   val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   839   val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   840   val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   841   val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   842   val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   843   val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   844   val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   845   val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   846   val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   847   val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   848   val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   849   val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   850   val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   851   val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   852   val (p,_,fb,nxt,_,pt) = me nxt p [] pt; 
   853   f2str fb;
   854 \<close>
   855 
   856 text\<open>\noindent We compare our results with the pre calculated upshot.\<close>
   857 
   858 ML \<open>
   859   if f2str fa = "[A = 4]" then () else error "part.fract. eq4_1";
   860   if f2str fb = "[B = -4]" then () else error "part.fract. eq4_1";
   861 \<close>
   862 
   863 section \<open>Implement the Specification and the MethodC \label{spec-meth}\<close>
   864 
   865 text\<open>\noindent Now everything we need for solving the problem has been
   866       tested out. We now start by creating new nodes for our methods and
   867       further on our new program in the interpreter.\<close>
   868 
   869 subsection\<open>Define the Field Descriptions for the 
   870             Specification\label{sec:deffdes}\<close>
   871 
   872 text\<open>\noindent We define the fields \em filterExpression \normalfont and
   873       \em stepResponse \normalfont both as equations, they represent the in- and
   874       output of the program.\<close>
   875 
   876 consts
   877   filterExpression  :: "bool => una"
   878   stepResponse      :: "bool => una"
   879 
   880 subsection\<open>Define the Specification\<close>
   881 
   882 text\<open>\noindent The next step is defining the specifications as nodes in the
   883       designated part. We have to create the hierarchy node by node and start
   884       with \em SignalProcessing \normalfont and go on by creating the node
   885       \em Z\_Transform\normalfont.\<close>
   886 
   887 setup \<open>KEStore_Elems.add_pbts
   888   [Problem.prep_input @{theory} "pbl_SP" [] Problem.id_empty (["SignalProcessing"], [], Rule_Set.empty, NONE, []),
   889     Problem.prep_input @{theory} "pbl_SP_Ztrans" [] Problem.id_empty
   890       (["Z_Transform", "SignalProcessing"], [], Rule_Set.empty, NONE, [])]\<close>
   891 
   892 text\<open>\noindent For the suddenly created node we have to define the input
   893        and output parameters. We already prepared their definition in
   894        Section~\ref{sec:deffdes}.\<close>
   895 
   896 setup \<open>KEStore_Elems.add_pbts
   897   [Problem.prep_input @{theory} "pbl_SP_Ztrans_inv" [] Problem.id_empty
   898       (["Inverse", "Z_Transform", "SignalProcessing"],
   899         [("#Given" ,["filterExpression X_eq"]),
   900           ("#Find", ["stepResponse n_eq"])],
   901         Rule_Set.append_rules "empty" Rule_Set.empty [(*for preds in where_*)],
   902         NONE,
   903         [["SignalProcessing", "Z_Transform", "Inverse"]])]\<close>
   904 ML \<open>
   905   Test_Tool.show_ptyps() ();
   906   Problem.from_store ["Inverse", "Z_Transform", "SignalProcessing"];
   907 \<close>
   908 
   909 subsection \<open>Define Name and Signature for the MethodC\<close>
   910 
   911 text\<open>\noindent As a next step we store the definition of our new method as a
   912       constant for the interpreter.\<close>
   913 
   914 consts
   915   InverseZTransform :: "[bool, bool] => bool"
   916     ("((Program InverseZTransform (_ =))// (_))" 9)
   917 
   918 subsection \<open>Setup Parent Nodes in Hierarchy of MethodC\label{sec:cparentnode}\<close>
   919 
   920 text\<open>\noindent Again we have to generate the nodes step by step, first the
   921       parent node and then the originally \em Z\_Transformation 
   922       \normalfont node. We have to define both nodes first with an empty script
   923       as content.\<close>
   924 
   925 setup \<open>KEStore_Elems.add_mets
   926   [MethodC.prep_input @{theory} "met_SP" [] e_metID
   927       (["SignalProcessing"], [],
   928         {rew_ord'="tless_true", rls'= Rule_Set.empty, calc = [], srls = Rule_Set.empty, prls = Rule_Set.empty, crls = Rule_Set.empty,
   929           errpats = [], nrls = Rule_Set.empty},
   930         "empty_script"),
   931     MethodC.prep_input @{theory} "met_SP_Ztrans" [] e_metID
   932       (["SignalProcessing", "Z_Transform"], [],
   933         {rew_ord'="tless_true", rls'= Rule_Set.empty, calc = [], srls = Rule_Set.empty, prls = Rule_Set.empty, crls = Rule_Set.empty,
   934           errpats = [], nrls = Rule_Set.empty},
   935         "empty_script")]
   936 \<close>
   937 
   938 text\<open>\noindent After we generated both nodes, we can fill the containing
   939       script we want to implement later. First we define the specifications
   940       of the script in e.g. the in- and output.\<close>
   941 
   942 setup \<open>KEStore_Elems.add_mets
   943   [MethodC.prep_input thy "met_SP_Ztrans_inv" [] e_metID
   944       (["SignalProcessing", "Z_Transform", "Inverse"], 
   945         [("#Given" ,["filterExpression X_eq", "boundVariable X_z"]),
   946           ("#Find", ["stepResponse n_eq"])],
   947         {rew_ord'="tless_true", rls'= Rule_Set.empty, calc = [], srls = Rule_Set.empty, prls = Rule_Set.empty, crls = Rule_Set.empty,
   948           errpats = [], nrls = Rule_Set.empty},
   949         "empty_script")]
   950 \<close>
   951 
   952 text\<open>\noindent After we stored the definition we can start implementing the
   953       script itself. As a first try we define only three rows containing one
   954       simple operation.\<close>
   955 
   956 setup \<open>KEStore_Elems.add_mets
   957   [MethodC.prep_input thy "met_SP_Ztrans_inv" [] e_metID
   958       (["SignalProcessing", "Z_Transform", "Inverse"], 
   959         [("#Given" , ["filterExpression X_eq", "boundVariable X_z"]),
   960           ("#Find", ["stepResponse n_eq"])],
   961         {rew_ord'="tless_true", rls'= Rule_Set.empty, calc = [], srls = Rule_Set.empty, prls = Rule_Set.empty, crls = Rule_Set.empty,
   962           errpats = [], nrls = Rule_Set.empty},
   963         "Program InverseZTransform (Xeq::bool) =" ^ (*TODO boundVariable X_z*)
   964           " (let X = Take Xeq;" ^
   965           "      X = Rewrite ruleZY X" ^
   966           "  in X)")]
   967 \<close>
   968 
   969 text\<open>\noindent Check if the method has been stored correctly\ldots\<close>
   970 
   971 ML \<open>
   972   show_mets(); 
   973 \<close>
   974 
   975 text\<open>\noindent If yes we can get the method by stepping backwards through
   976       the hierarchy.\<close>
   977 
   978 ML \<open>
   979   MethodC.from_store ["SignalProcessing", "Z_Transform", "Inverse"];
   980 \<close>
   981 
   982 section \<open>Program in TP-based language \label{prog-steps}\<close>
   983 
   984 text\<open>\noindent We start stepwise expanding our program. The script is a
   985       simple string containing several manipulation instructions.
   986       \par The first script we try contains no instruction, we only test if
   987       building scripts that way work.\<close>
   988 
   989 subsection \<open>Stepwise Extend the Program\<close>
   990 ML \<open>
   991   val str = 
   992     "Program InverseZTransform (Xeq::bool) =                          "^
   993     " Xeq";
   994 \<close>
   995 
   996 text\<open>\noindent Next we put some instructions in the script and try if we are
   997       able to solve our first equation.\<close>
   998 
   999 ML \<open>
  1000   val str = 
  1001     "Program InverseZTransform (Xeq::bool) =                          "^
  1002     (*
  1003      * 1/z) instead of z \<up> -1
  1004      *)
  1005     " (let X = Take Xeq;                                             "^
  1006     "      X' = Rewrite ruleZY False X;                              "^
  1007     (*
  1008      * z * denominator
  1009      *)
  1010     "      X' = (Rewrite_Set norm_Rational False) X'                 "^
  1011     (*
  1012      * simplify
  1013      *)
  1014     "  in X)";
  1015     (*
  1016      * NONE
  1017      *)
  1018     "Program InverseZTransform (Xeq::bool) =                          "^
  1019     (*
  1020      * (1/z) instead of z \<up> -1
  1021      *)
  1022     " (let X = Take Xeq;                                             "^
  1023     "      X' = Rewrite ruleZY False X;                              "^
  1024     (*
  1025      * z * denominator
  1026      *)
  1027     "      X' = (Rewrite_Set norm_Rational False) X';                "^
  1028     (*
  1029      * simplify
  1030      *)
  1031     "      X' = (SubProblem (IsacX,[pqFormula,degree_2,              "^
  1032     "                               polynomial,univariate,equation], "^
  1033     "                              [no_met])                         "^
  1034     "                              [BOOL e_e, REAL v_v])             "^
  1035     "            in X)";
  1036 \<close>
  1037 
  1038 text\<open>\noindent To solve the equation it is necessary to drop the left hand
  1039       side, now we only need the denominator of the right hand side. The first
  1040       equation solves the zeros of our expression.\<close>
  1041 
  1042 ML \<open>
  1043   val str = 
  1044     "Program InverseZTransform (Xeq::bool) =                          "^
  1045     " (let X = Take Xeq;                                             "^
  1046     "      X' = Rewrite ruleZY False X;                              "^
  1047     "      X' = (Rewrite_Set norm_Rational False) X';                "^
  1048     "      funterm = rhs X'                                          "^
  1049     (*
  1050      * drop X'= for equation solving
  1051      *)
  1052     "  in X)";
  1053 \<close>
  1054 
  1055 text\<open>\noindent As mentioned above, we need the denominator of the right hand
  1056       side. The equation itself consists of this denominator and tries to find
  1057       a $z$ for this the denominator is equal to zero.\<close>
  1058 
  1059 text \<open> dropped during switch from Program to partial_function:
  1060   val str = 
  1061     "Program InverseZTransform (X_eq::bool) =                         "^
  1062     " (let X = Take X_eq;                                            "^
  1063     "      X' = Rewrite ruleZY False X;                              "^
  1064     "      X' = (Rewrite_Set norm_Rational False) X';                "^
  1065     "      (X'_z::real) = lhs X';                                    "^
  1066     "      (z::real) = argument_in X'_z;                             "^
  1067     "      (funterm::real) = rhs X';                                 "^
  1068     "      (denom::real) = get_denominator funterm;                  "^
  1069     (*
  1070      * get_denominator
  1071      *)
  1072     "      (equ::bool) = (denom = (0::real));                        "^
  1073     "      (L_L::bool list) =                                        "^
  1074     "            (SubProblem (Test,                                 "^
  1075     "                         [LINEAR,univariate,equation,test],     "^
  1076     "                         [Test,solve_linear])                   "^
  1077     "                         [BOOL equ, REAL z])                    "^
  1078     "  in X)";
  1079 
  1080   parse thy str;
  1081   val sc = (inst_abs o Thm.term_of o the o (parse thy)) str;
  1082   atomty sc;
  1083 \<close>
  1084 
  1085 text \<open>\noindent This ruleset contains all functions that are in the script; 
  1086        The evaluation of the functions is done by rewriting using this ruleset.\<close>
  1087 
  1088 ML \<open>
  1089   val srls = 
  1090     Rule_Set.Repeat {id="srls_InverseZTransform", 
  1091          preconds = [],
  1092          rew_ord = ("termlessI",termlessI),
  1093          erls = Rule_Set.append_rules "erls_in_srls_InverseZTransform" Rule_Set.empty
  1094            [(*for asm in NTH_CONS ...*)
  1095             Eval ("Orderings.ord_class.less",eval_equ "#less_"),
  1096             (*2nd NTH_CONS pushes n+-1 into asms*)
  1097             Eval("Groups.plus_class.plus", eval_binop "#add_")
  1098            ], 
  1099          srls = Rule_Set.Empty, calc = [], errpatts = [],
  1100          rules = [
  1101                   Thm ("NTH_CONS",ThmC.numerals_to_Free @{thm NTH_CONS}),
  1102                   Eval("Groups.plus_class.plus", eval_binop "#add_"),
  1103                   Thm ("NTH_NIL",ThmC.numerals_to_Free @{thm NTH_NIL}),
  1104                   Eval("Prog_Expr.lhs", eval_lhs"eval_lhs_"),
  1105                   Eval("Prog_Expr.rhs", eval_rhs"eval_rhs_"),
  1106                   Eval("Prog_Expr.argument'_in", eval_argument_in "Prog_Expr.argument'_in"),
  1107                   Eval("Rational.get_denominator", eval_get_denominator "#get_denominator"),
  1108                   Eval("Rational.get_numerator", eval_get_numerator "#get_numerator"),
  1109                   Eval("Partial_Fractions.factors_from_solution",
  1110                        eval_factors_from_solution "#factors_from_solution"),
  1111                   Eval("Partial_Fractions.drop_questionmarks",
  1112                        eval_drop_questionmarks "#drop_?")
  1113                  ],
  1114          scr = Empty_Prog};
  1115 \<close>
  1116 
  1117 
  1118 subsection \<open>Store Final Version of Program for Execution\<close>
  1119 
  1120 text\<open>\noindent After we also tested how to write scripts and run them,
  1121       we start creating the final version of our script and store it into
  1122       the method for which we created a node in Section~\ref{sec:cparentnode}
  1123       Note that we also did this stepwise, but we didn't kept every
  1124       subversion.\<close>
  1125 
  1126 setup \<open>KEStore_Elems.add_mets
  1127   [MethodC.prep_input thy "met_SP_Ztrans_inv" [] e_metID
  1128       (["SignalProcessing", "Z_Transform", "Inverse"], 
  1129         [("#Given" , ["filterExpression X_eq"]), (*TODO boundVariable X_z*)
  1130           ("#Find", ["stepResponse n_eq"])],
  1131         {rew_ord'="tless_true", rls'= Rule_Set.empty, calc = [], srls = srls, prls = Rule_Set.empty, crls = Rule_Set.empty,
  1132           errpats = [], nrls = Rule_Set.empty},
  1133         "Program InverseZTransform (X_eq::bool) =                        "^
  1134           (*(1/z) instead of z \<up> -1*)
  1135           "(let X = Take X_eq;                                            "^
  1136           "      X' = Rewrite ruleZY False X;                             "^
  1137           (*z * denominator*)
  1138           "      (num_orig::real) = get_numerator (rhs X');               "^
  1139           "      X' = (Rewrite_Set norm_Rational False) X';               "^
  1140           (*simplify*)
  1141           "      (X'_z::real) = lhs X';                                   "^
  1142           "      (zzz::real) = argument_in X'_z;                          "^
  1143           "      (funterm::real) = rhs X';                                "^
  1144           (*drop X' z = for equation solving*)
  1145           "      (denom::real) = get_denominator funterm;                 "^
  1146           (*get_denominator*)
  1147           "      (num::real) = get_numerator funterm;                     "^
  1148           (*get_numerator*)
  1149           "      (equ::bool) = (denom = (0::real));                       "^
  1150           "      (L_L::bool list) = (SubProblem (PolyEqX,                 "^
  1151           "         [abcFormula,degree_2,polynomial,univariate,equation], "^
  1152           "         [no_met])                                             "^
  1153           "         [BOOL equ, REAL zzz]);                                "^
  1154           "      (facs::real) = factors_from_solution L_L;                "^
  1155           "      (eql::real) = Take (num_orig / facs);                    "^ 
  1156 
  1157           "      (eqr::real) = (Try (Rewrite_Set ansatz_rls False)) eql;  "^
  1158 
  1159           "      (eq::bool) = Take (eql = eqr);                           "^
  1160           (*Maybe possible to use HOLogic.mk_eq ??*)
  1161           "      eq = (Try (Rewrite_Set equival_trans False)) eq;         "^ 
  1162 
  1163           "      eq = drop_questionmarks eq;                              "^
  1164           "      (z1::real) = (rhs (NTH 1 L_L));                          "^
  1165           (* 
  1166           * prepare equation for a - eq_a
  1167           * therefor substitute z with solution 1 - z1
  1168           *)
  1169           "      (z2::real) = (rhs (NTH 2 L_L));                          "^
  1170 
  1171           "      (eq_a::bool) = Take eq;                                  "^
  1172           "      eq_a = (Substitute [zzz=z1]) eq;                         "^
  1173           "      eq_a = (Rewrite_Set norm_Rational False) eq_a;           "^
  1174           "      (sol_a::bool list) =                                     "^
  1175           "                 (SubProblem (IsacX,                           "^
  1176           "                              [univariate,equation],[no_met])  "^
  1177           "                              [BOOL eq_a, REAL (A::real)]);    "^
  1178           "      (a::real) = (rhs(NTH 1 sol_a));                          "^
  1179 
  1180           "      (eq_b::bool) = Take eq;                                  "^
  1181           "      eq_b =  (Substitute [zzz=z2]) eq_b;                      "^
  1182           "      eq_b = (Rewrite_Set norm_Rational False) eq_b;           "^
  1183           "      (sol_b::bool list) =                                     "^
  1184           "                 (SubProblem (IsacX,                           "^
  1185           "                              [univariate,equation],[no_met])  "^
  1186           "                              [BOOL eq_b, REAL (B::real)]);    "^
  1187           "      (b::real) = (rhs(NTH 1 sol_b));                          "^
  1188 
  1189           "      eqr = drop_questionmarks eqr;                            "^
  1190           "      (pbz::real) = Take eqr;                                  "^
  1191           "      pbz = ((Substitute [A=a, B=b]) pbz);                     "^
  1192 
  1193           "      pbz = Rewrite ruleYZ False pbz;                          "^
  1194           "      pbz = drop_questionmarks pbz;                            "^
  1195 
  1196           "      (X_z::bool) = Take (X_z = pbz);                          "^
  1197           "      (n_eq::bool) = (Rewrite_Set inverse_z False) X_z;        "^
  1198           "      n_eq = drop_questionmarks n_eq                           "^
  1199           "in n_eq)")]
  1200 \<close>
  1201 
  1202 
  1203 subsection \<open>Check the Program\<close>
  1204 text\<open>\noindent When the script is ready we can start checking our work.\<close>
  1205 subsubsection \<open>Check the Formalization\<close>
  1206 text\<open>\noindent First we want to check the formalization of the in and
  1207        output of our program.\<close>
  1208 
  1209 ML \<open>
  1210   val fmz = 
  1211     ["filterExpression (X  = 3 / (z - 1/4 + -1/8 * (1/(z::real))))",
  1212      "stepResponse (x[n::real]::bool)"];
  1213   val (dI,pI,mI) = 
  1214     ("Isac_Knowledge", ["Inverse", "Z_Transform", "SignalProcessing"], 
  1215              ["SignalProcessing", "Z_Transform", "Inverse"]);
  1216 
  1217   val ([
  1218           (
  1219             1,
  1220             [1],
  1221             "#Given",
  1222             Const ("Inverse_Z_Transform.filterExpression", _),
  1223             [Const ("HOL.eq", _) $ _ $ _]
  1224           ),
  1225           (
  1226             2,
  1227             [1],
  1228             "#Find",
  1229             Const ("Inverse_Z_Transform.stepResponse", _),
  1230             [Free ("x", _) $ _]
  1231           )
  1232        ],_
  1233       ) = O_Model.init fmz thy ((#ppc o Problem.from_store) pI);
  1234 
  1235   val Prog sc 
  1236     = (#scr o MethodC.from_store) ["SignalProcessing",
  1237                         "Z_Transform",
  1238                         "Inverse"];
  1239   atomty sc;
  1240 \<close>
  1241 
  1242 subsubsection \<open>Stepwise Check the Program\label{sec:stepcheck}\<close>
  1243 text\<open>\noindent We start to stepwise execute our new program in a calculation
  1244       tree and check if every node implements that what we have wanted.\<close>
  1245       
  1246 ML \<open>
  1247   Rewrite.trace_on := false; (*true*)
  1248   LItool.trace_on := false; (*true*)
  1249   print_depth 9;
  1250   
  1251   val fmz =
  1252     ["filterExpression (X z = 3 / ((z::real) + 1/10 - 1/50*(1/z)))",
  1253      "stepResponse (x[n::real]::bool)"];
  1254      
  1255   val (dI,pI,mI) =
  1256     ("Isac_Knowledge", ["Inverse", "Z_Transform", "SignalProcessing"], 
  1257              ["SignalProcessing", "Z_Transform", "Inverse"]);
  1258              
  1259   val (p,_,f,nxt,_,pt)  = CalcTreeTEST [(fmz, (dI,pI,mI))];
  1260   val (p,_,f,nxt,_,pt) = me nxt p [] pt; 
  1261     "Add_Given";
  1262   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1263     "Add_Find";
  1264   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1265     "Specify_Theory";
  1266   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1267     "Specify_Problem";
  1268   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1269     "Specify_Method";
  1270   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1271     "Apply_Method";
  1272   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1273     "Rewrite (ruleZY, Inverse_Z_Transform.ruleZY)";
  1274     "--> X z = 3 / (z - 1 / 4 + -1 / 8 * (1 / z))";
  1275   (*
  1276    * TODO naming!
  1277    *)
  1278   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1279     "Rewrite_Set norm_Rational";
  1280     " --> X' z = 3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z)))";
  1281   val (p,_,f,nxt,_,pt) = me nxt p [] pt; 
  1282     "SubProblem";
  1283   print_depth 3;
  1284 \<close>
  1285 
  1286 text \<open>\noindent Instead of \ttfamily nxt = Subproblem \normalfont above we had
  1287        \ttfamily Empty\_Tac; \normalfont the search for the reason considered
  1288        the following points:\begin{itemize}
  1289        \item What shows \ttfamily show\_pt pt;\normalfont\ldots?
  1290 \begin{verbatim}(([2], Res), ?X' z = 24 / (-1 + -2 * z + 8 * z \<up> 2))]\end{verbatim}
  1291          The calculation is ok but no \ttfamily next \normalfont step found:
  1292          Should be\\ \ttfamily nxt = Subproblem\normalfont!
  1293        \item What shows \ttfamily trace\_script := true; \normalfont we read
  1294          bottom up\ldots
  1295      \begin{verbatim}
  1296      @@@next leaf 'SubProblem
  1297      (PolyEq',[abcFormula, degree_2, polynomial, 
  1298                univariate, equation], no_meth)
  1299      [BOOL equ, REAL z]' 
  1300        ---> Program.Tac 'SubProblem (PolyEq',
  1301               [abcFormula, degree_2, polynomial,
  1302                univariate, equation], no_meth)
  1303      [BOOL (-1 + -2 * z + 8 * z \^\^\^ ~2 = 0), REAL z]'
  1304      \end{verbatim}
  1305          We see the SubProblem with correct arguments from searching next
  1306          step (program text !!!--->!!! Program.Tac (script tactic) with arguments
  1307          evaluated.)
  1308      \item Do we have the right Program \ldots difference in the
  1309          arguments in the arguments\ldots
  1310          \begin{verbatim}
  1311      val Prog s =
  1312      (#scr o MethodC.from_store) ["SignalProcessing",
  1313                        "Z_Transform",
  1314                        "Inverse"];
  1315      writeln (UnparseC.term s);
  1316          \end{verbatim}
  1317          \ldots shows the right script. Difference in the arguments.
  1318      \item Test --- Why helpless here ? --- shows: \ttfamily replace
  1319          no\_meth by [no\_meth] \normalfont in Program
  1320      \end{itemize}
  1321 \<close>
  1322 
  1323 ML \<open>
  1324   val (p,_,f,nxt,_,pt) = me nxt p [] pt; 
  1325     (*Model_Problem";*)
  1326 \<close>
  1327 
  1328 text \<open>\noindent Instead of \ttfamily nxt = Model\_Problem \normalfont above
  1329        we had \ttfamily Empty\_Tac; \normalfont the search for the reason 
  1330        considered the following points:\begin{itemize}
  1331        \item Difference in the arguments
  1332        \item Comparison with Subsection~\ref{sec:solveq}: There solving this
  1333          equation works, so there must be some difference in the arguments
  1334          of the Subproblem: RIGHT: we had \ttfamily [no\_meth] \normalfont
  1335          instead of \ttfamily [no\_met] \normalfont ;-)
  1336        \end{itemize}\<close>
  1337 
  1338 ML \<open>
  1339   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1340     (*Add_Given equality (-1 + -2 * z + 8 * z \<up> 2 = 0)";*)
  1341   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1342     (*Add_Given solveFor z";*)
  1343   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1344     (*Add_Find solutions z_i";*)
  1345   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1346     (*Specify_Theory Isac";*)
  1347 \<close>
  1348 
  1349 text \<open>\noindent We had \ttfamily nxt = Empty\_Tac instead Specify\_Theory;
  1350        \normalfont The search for the reason considered the following points:
  1351        \begin{itemize}
  1352        \item Was there an error message? NO -- ok
  1353        \item Has \ttfamily nxt = Add\_Find \normalfont been inserted in pt:\\
  1354          \ttfamily get\_obj g\_pbl pt (fst p);\normalfont? YES -- ok
  1355        \item What is the returned formula:
  1356 \begin{verbatim}
  1357 print_depth 999; f; print_depth 3;
  1358 { Find = [ Correct "solutions z_i"],
  1359   With = [],
  1360   Given = [Correct "equality (-1 + -2*z + 8*z \<up> 2 = 0)",
  1361            Correct "solveFor z"],
  1362   Where = [...],
  1363   Relate = [] }
  1364 \end{verbatim}
  1365      \normalfont The only False is the reason: the Where (the precondition) is
  1366      False for good reasons: The precondition seems to check for linear
  1367      equations, not for the one we want to solve! Removed this error by
  1368      correcting the Program from \ttfamily SubProblem (PolyEq',
  1369      \lbrack linear,univariate,equation,
  1370        test\rbrack, \lbrack Test,solve\_linear\rbrack \normalfont to
  1371      \ttfamily SubProblem (PolyEq',\\ \lbrack abcFormula,degree\_2,
  1372        polynomial,univariate,equation\rbrack,\\
  1373                    \lbrack PolyEq,solve\_d2\_polyeq\_abc\_equation
  1374                    \rbrack\normalfont
  1375      You find the appropriate type of equation at the
  1376      {\sisac}-WEB-Page\footnote{
  1377      \href{http://www.ist.tugraz.at/projects/isac/www/kbase/pbl/index\_pbl.html}
  1378           {http://www.ist.tugraz.at/projects/isac/www/kbase/pbl/index\_pbl.html}
  1379                                }
  1380      And the respective method in \ttfamily Knowledge/PolyEq.thy \normalfont
  1381      at the respective \ttfamily store\_pbt. \normalfont Or you leave the
  1382      selection of the appropriate type to isac as done in the final Program ;-))
  1383   \end{itemize}\<close>
  1384   
  1385 ML \<open>
  1386   val (p,_,f,nxt,_,pt) = me nxt p [] pt; 
  1387     (*Specify_Problem [abcFormula,...";*)
  1388   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1389     (*Specify_Method [PolyEq,solve_d2_polyeq_abc_equation";*)
  1390   val (p,_,f,nxt,_,pt) = me nxt p [] pt; 
  1391     (*Apply_Method [PolyEq,solve_d2_polyeq_abc_equation";*)
  1392   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1393     (*Rewrite_Set_Inst [(''bdv'', z)], d2_polyeq_abcFormula_simplify";*)
  1394   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1395   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1396   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1397   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1398   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1399   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1400   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1401   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1402   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1403   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1404   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1405   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1406   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1407   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1408   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1409   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1410   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1411     (*Specify_Problem ["normalise", "polynomial", "univariate", "equation"]*)
  1412   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1413     (*Specify_Method ["PolyEq", "normalise_poly"]*)
  1414   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1415     (*Apply_Method ["PolyEq", "normalise_poly"]*)
  1416   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1417     (*Rewrite ("all_left", "PolyEq.all_left")*)
  1418   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1419     (*Rewrite_Set_Inst (["(''bdv'', A)"], "make_ratpoly_in")*)
  1420   val (p,_,f,nxt,_,pt) = me nxt p [] pt; 
  1421     (*Rewrite_Set "polyeq_simplify"*)
  1422   val (p,_,f,nxt,_,pt) = me nxt p [] pt; 
  1423     (*Subproblem("Isac_Knowledge",["degree_1", "polynomial", "univariate", "equation"])*)
  1424   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1425   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1426   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1427   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1428   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1429   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1430   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1431   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1432   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1433   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1434   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1435   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1436   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1437   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1438   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1439   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1440   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1441   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1442   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1443   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1444   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1445   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1446   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1447   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1448   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1449   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1450   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1451   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1452   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1453   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1454   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1455   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1456   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1457   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1458   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1459   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1460   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1461   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1462   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1463   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1464   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1465   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1466   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1467   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1468   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1469   val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*([11, 4, 5], Res) Check_Postcond*)
  1470   val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*([11, 4], Res) Check_Postcond*)
  1471   val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*([11], Res) Take*)
  1472   val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*([12], Frm) Substitute*)
  1473   val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*([12], Res) Rewrite*)
  1474   val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*([13], Res) Take*)
  1475   val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*([14], Frm) Empty_Tac*)
  1476 \<close>
  1477 ML \<open>
  1478   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1479 \<close>
  1480 ML \<open>
  1481   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1482 \<close>
  1483 ML \<open>
  1484   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1485 \<close>
  1486 ML \<open>
  1487   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1488 \<close>
  1489 ML \<open>
  1490   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1491 \<close>
  1492 ML \<open>
  1493   val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1494 \<close>
  1495 
  1496 ML \<open>
  1497 LItool.trace_on := true;
  1498 \<close>
  1499 ML \<open>
  1500 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1501 \<close>
  1502 ML \<open>
  1503 Test_Tool.show_pt pt;
  1504 \<close>
  1505 ML \<open>
  1506 \<close> 
  1507 ML \<open>
  1508 \<close> 
  1509 ML \<open>
  1510 \<close> 
  1511 ML \<open>
  1512 \<close> 
  1513 ML \<open>
  1514 \<close> 
  1515 ML \<open>
  1516 \<close> 
  1517 
  1518 text\<open>\noindent As a last step we check the tracing output of the last calc
  1519       tree instruction and compare it with the pre-calculated result.\<close>
  1520 
  1521 section \<open>Improve and Transfer into Knowledge\<close>
  1522 text \<open>
  1523   We want to improve the very long program \ttfamily InverseZTransform
  1524   \normalfont by modularisation: partial fraction decomposition shall
  1525   become a sub-problem.
  1526 
  1527   We could transfer all knowledge in \ttfamily Build\_Inverse\_Z\_Transform.thy 
  1528   \normalfont first to the \ttfamily Knowledge/Inverse\_Z\_Transform.thy 
  1529   \normalfont and then modularise. In this case TODO problems?!?
  1530 
  1531   We chose another way and go bottom up: first we build the sub-problem in
  1532   \ttfamily Partial\_Fractions.thy \normalfont with the term:
  1533 
  1534       $$\frac{3}{x\cdot(z - \frac{1}{4} + \frac{-1}{8}\cdot\frac{1}{z})}$$
  1535 
  1536   \noindent (how this still can be improved see \ttfamily Partial\_Fractions.thy\normalfont),
  1537   and re-use all stuff prepared in \ttfamily Build\_Inverse\_Z\_Transform.thy:
  1538   \normalfont The knowledge will be transferred to \ttfamily src/../Partial\_Fractions.thy 
  1539   \normalfont and the respective tests to:
  1540   \begin{center}\ttfamily test/../sartial\_fractions.sml\normalfont\end{center}
  1541 \<close>
  1542 
  1543 subsection \<open>Transfer to Partial\_Fractions.thy\<close>
  1544 text \<open>
  1545   First we transfer both, knowledge and tests into:
  1546   \begin{center}\ttfamily src/../Partial\_Fractions.thy\normalfont\end{center}
  1547   in order to immediately have the test results.
  1548 
  1549   We copy \ttfamily factors\_from\_solution, drop\_questionmarks,\\
  1550   ansatz\_2nd\_order \normalfont and rule-sets --- no problem.
  1551   
  1552   Also \ttfamily store\_pbt ..\\ "pbl\_simp\_rat\_partfrac"
  1553   \normalfont is easy.
  1554 
  1555   But then we copy from:\\
  1556   (1) \ttfamily Build\_Inverse\_Z\_Transform.thy store\_met\ldots "met\_SP\_Ztrans\_inv"
  1557   \normalfont\\ to\\ 
  1558   (2) \ttfamily Partial\_Fractions.thy store\_met\ldots "met\_SP\_Ztrans\_inv" 
  1559   \normalfont\\ and cut out the respective part from the program. First we ensure that
  1560   the string is correct. When we insert the string into (2)
  1561   \ttfamily store\_met .. "met\_partial\_fraction" \normalfont --- and get an error.
  1562 \<close>
  1563 
  1564 subsubsection \<open>'Programming' in ISAC's TP-based Language\<close>
  1565 text \<open>
  1566   At the present state writing programs in {\sisac} is particularly cumbersome.
  1567   So we give hints how to cope with the many obstacles. Below we describe the
  1568   steps we did in making (2) run.
  1569   
  1570   \begin{enumerate}
  1571     \item We check if the \textbf{string} containing the program is correct.
  1572     \item We check if the \textbf{types in the program} are correct.
  1573       For this purpose we start start with the first and last lines
  1574      \begin{verbatim}
  1575      "PartFracScript (f_f::real) (v_v::real) =       " ^
  1576      " (let X = Take f_f;                            " ^
  1577      "      pbz = ((Substitute []) X)                " ^
  1578      "  in pbz)"
  1579      \end{verbatim}
  1580        The last but one line helps not to bother with ';'.
  1581      \item Then we add line by line. Already the first line causes the error. 
  1582         So we investigate it by
  1583       \begin{verbatim}
  1584       val ctxt = Proof_Context.init_global @{theory "Inverse_Z_Transform"} ;
  1585       val SOME t = 
  1586         parseNEW ctxt "(num_orig::real) = 
  1587                           get_numerator(rhs f_f)";
  1588       \end{verbatim}
  1589         and see a type clash: \ttfamily rhs \normalfont from (1) requires type 
  1590         \ttfamily bool \normalfont while (2) wants to have \ttfamily (f\_f::real).
  1591         \normalfont Of course, we don't need \ttfamily rhs \normalfont anymore.
  1592       \item Type-checking can be very tedious. One might even inspect the
  1593         parse-tree of the program with {\sisac}'s specific debug tools:
  1594       \begin{verbatim}
  1595       val {scr = Prog t,...} = 
  1596         MethodC.from_store ["simplification",
  1597                  "of_rationals",
  1598                  "to_partial_fraction"];
  1599       atomty_thy @{theory "Inverse_Z_Transform"} t ;
  1600       \end{verbatim}
  1601       \item We check if the \textbf{semantics of the program} by stepwise evaluation
  1602         of the program. Evaluation is done by the Lucas-Interpreter, which works
  1603         using the knowledge in theory Isac_Knowledge; so we have to re-build Isac. And the
  1604         test are performed simplest in a file which is loaded with Isac.
  1605         See \ttfamily tests/../partial\_fractions.sml \normalfont.
  1606   \end{enumerate}
  1607 \<close>
  1608 
  1609 subsection \<open>Transfer to Inverse\_Z\_Transform.thy\<close>
  1610 text \<open>
  1611   It was not possible to complete this task, because we ran out of time.
  1612 \<close>
  1613 
  1614 
  1615 end
  1616