test/Tools/isac/ADDTESTS/course/SignalProcess/Build_Inverse_Z_Transform.thy
author Jan Rocnik <jan.rocnik@student.tugraz.at>
Tue, 14 Feb 2012 22:55:03 +0100
changeset 42368 3afe632cd527
parent 42367 c1ebb7e759f9
child 42369 3fa947c99a28
permissions -rwxr-xr-x
tuned - partitially formated
     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 \ttfamily Inverse\_Z\_Transform.thy \normalfont as an 
    13   exercise. Because subsection~\ref{sec:stepcheck} requires 
    14   \ttfamily Inverse\_Z\_Transform.thy \normalfont as a subtheory of \ttfamily 
    15   Isac.thy\normalfont, the setup has been changed from \ttfamily theory 
    16   Inverse\_Z\_Transform imports Isac \normalfont to the above one.
    17   \par \noindent
    18   \begin{center} 
    19   \textbf{ATTENTION WITH NAMES OF IDENTIFIERS WHEN GOING INTO INTERNALS}
    20   \end{center}
    21   Here in this theory there are the internal names twice, for instance we have
    22   \ttfamily (Thm.derivation\_name @{thm rule1} = 
    23   "Build\_Inverse\_Z\_Transform.rule1") = true; \normalfont
    24   but actually in us will be \ttfamily Inverse\_Z\_Transform.rule1 \normalfont
    25 *}
    26 
    27 section {*Trials towards the Z-Transform\label{sec:trials}*}
    28 
    29 ML {*val thy = @{theory Isac};*}
    30 
    31 subsection {*Notations and Terms*}
    32 text{*\noindent Try which notations we are able to use.*}
    33 ML {*
    34   @{term "1 < || z ||"};
    35   @{term "z / (z - 1)"};
    36   @{term "-u -n - 1"};
    37   @{term "-u [-n - 1]"}; (*[ ] denotes lists !!!*)
    38   @{term "z /(z - 1) = -u [-n - 1]"};Isac
    39   @{term "1 < || z || ==> z / (z - 1) = -u [-n - 1]"};
    40   term2str @{term "1 < || z || ==> z / (z - 1) = -u [-n - 1]"};
    41 *}
    42 text{*\noindent Try which symbols we are able to use and how we generate them.*}
    43 ML {*
    44   (*alpha -->  "</alpha>" *)
    45   @{term "\<alpha> "};
    46   @{term "\<delta> "};
    47   @{term "\<phi> "};
    48   @{term "\<rho> "};
    49   term2str @{term "\<rho> "};
    50 *}
    51 
    52 subsection {*Rules*}
    53 (*axiomatization "z / (z - 1) = -u [-n - 1]"
    54   Illegal variable name: "z / (z - 1) = -u [-n - 1]" *)
    55 (*definition     "z / (z - 1) = -u [-n - 1]"
    56   Bad head of lhs: existing constant "op /"*)
    57 axiomatization where 
    58   rule1: "1 = \<delta>[n]" and
    59   rule2: "|| z || > 1 ==> z / (z - 1) = u [n]" and
    60   rule3: "|| z || < 1 ==> z / (z - 1) = -u [-n - 1]" and 
    61   rule4: "|| z || > || \<alpha> || ==> z / (z - \<alpha>) = \<alpha>^^^n * u [n]" and
    62   rule5: "|| z || < || \<alpha> || ==> z / (z - \<alpha>) = -(\<alpha>^^^n) * u [-n - 1]" and
    63   rule6: "|| z || > 1 ==> z/(z - 1)^^^2 = n * u [n]"
    64 
    65 text{*\noindent Check the rules for their correct notation. 
    66       (See the machine output.)*}
    67 ML {*
    68   @{thm rule1};
    69   @{thm rule2};
    70   @{thm rule3};
    71   @{thm rule4};
    72 *}
    73 
    74 subsection {*Apply Rules*}
    75 text{*\noindent We try to apply the rules to a given expression.*}
    76 
    77 ML {*
    78   val inverse_Z = append_rls "inverse_Z" e_rls
    79     [ Thm  ("rule3",num_str @{thm rule3}),
    80       Thm  ("rule4",num_str @{thm rule4}),
    81       Thm  ("rule1",num_str @{thm rule1})   
    82     ];
    83 
    84   val t = str2term "z / (z - 1) + z / (z - \<alpha>) + 1";
    85   val SOME (t', asm) = rewrite_set_ thy true inverse_Z t;
    86   term2str 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 *}
    92 
    93 ML {* val (thy, ro, er) = (@{theory Isac}, tless_true, eval_rls); *}
    94 ML {*
    95   val SOME (t, asm1) = 
    96     rewrite_ thy ro er true (num_str @{thm rule3}) t;
    97   term2str t = "- ?u [- ?n - 1] + z / (z - \<alpha>) + 1";
    98   (*- real *)
    99   term2str t;
   100 
   101   val SOME (t, asm2) = 
   102     rewrite_ thy ro er true (num_str @{thm rule4}) t;
   103   term2str t = "- ?u [- ?n - 1] + \<alpha> ^^^ ?n * ?u [?n] + 1";
   104   (*- real *)
   105   term2str t;
   106 
   107   val SOME (t, asm3) = 
   108     rewrite_ thy ro er true (num_str @{thm rule1}) t;
   109   term2str t = "- ?u [- ?n - 1] + \<alpha> ^^^ ?n * ?u [?n] + ?\<delta> [?n]";
   110   (*- real *)
   111   term2str t;
   112 *}
   113 ML {* terms2str (asm1 @ asm2 @ asm3); *}
   114 
   115 section{*Prepare Steps for CTP-based programming Language\label{sec:prepstep}*}
   116 text{*
   117       \par \noindent The following sections are challanging with the CTP-based 
   118       possibilities of building the programm. The goal is realized in 
   119       Section~\ref{spec-meth} and Section~\ref{prog-steps}.
   120       \par The reader is advised to jump between the subsequent subsections and 
   121       the respective steps in Section~\ref{prog-steps}. By comparing 
   122       Section~\ref{sec:calc:ztrans} the calculation can be comprehended step 
   123       by step.
   124 *}
   125 
   126 subsection {*Prepare Expression\label{prep-expr}*}
   127 ML {*
   128   val ctxt = ProofContext.init_global @{theory Isac};
   129   val ctxt = declare_constraints' [@{term "z::real"}] ctxt;
   130 
   131   val SOME fun1 = 
   132     parseNEW ctxt "X z = 3 / (z - 1/4 + -1/8 * z ^^^ -1)"; term2str fun1;
   133   val SOME fun1' = 
   134     parseNEW ctxt "X z = 3 / (z - 1/4 + -1/8 * (1/z))"; term2str fun1';
   135 *}
   136 
   137 subsubsection {*Prepare Numerator and Denominator*}
   138 text{*\noindent The partial fraction decomposion is only possible ig we
   139        get the bound variable out of the numerator. Therefor we divide
   140        the expression by $z$. Follow up the Calculation at 
   141        Section~\ref{sec:calc:ztrans} line number 02.*}
   142 
   143 axiomatization where
   144   ruleZY: "(X z = a / b) = (X' z = a / (z * b))"
   145 
   146 ML {*
   147   val (thy, ro, er) = (@{theory Isac}, tless_true, eval_rls);
   148   val SOME (fun2, asm1) = 
   149     rewrite_ thy ro er true  @{thm ruleZY} fun1; term2str fun2;
   150   val SOME (fun2', asm1) = 
   151     rewrite_ thy ro er true  @{thm ruleZY} fun1'; term2str fun2';
   152 
   153   val SOME (fun3,_) = 
   154     rewrite_set_ @{theory Isac} false norm_Rational fun2;
   155   term2str fun3;
   156   (*
   157    * Fails on x^^^(-1)
   158    * We solve this problem by using 1/x as a workaround.
   159    *)
   160   val SOME (fun3',_) = 
   161     rewrite_set_ @{theory Isac} false norm_Rational fun2';
   162   term2str fun3';
   163   (*
   164    * OK - workaround!
   165    *)
   166 *}
   167 
   168 subsubsection {*Get the Argument of the Expression X'*}
   169 text{*\noindent We use \texttt{grep} for finding possibilities how we can
   170        extract the bound variable in the expression. \ttfamily Atools.thy, 
   171        Tools.thy \normalfont contain general utilities: \ttfamily 
   172        eval\_argument\_in, eval\_rhs, eval\_lhs,\ldots \normalfont
   173        \ttfamily grep -r "fun eva\_" * \normalfont shows all functions 
   174        witch can be used in a script. Lookup this files how to build 
   175        and handle such functions.
   176        \par The next section shows how to introduce such a function.
   177 *}
   178 
   179 subsubsection{*Decompose the Given Term Into lhs and rhs\footnote{Note:
   180                lhs means \em Left Hand Side \normalfont and rhs means 
   181                \em Right Hand Side \normalfont and indicates the left or 
   182                the right part of an equation.}*}
   183 ML {*
   184   val (_, expr) = HOLogic.dest_eq fun3'; term2str expr;
   185   val (_, denom) = 
   186     HOLogic.dest_bin "Rings.inverse_class.divide" (type_of expr) expr;
   187   term2str denom = "-1 + -2 * z + 8 * z ^^^ 2";
   188 *}
   189 
   190 text{*\noindent We have rhs in the Script language, but we need a function 
   191        which gets the denominator of a fraction.*}
   192 
   193 subsubsection{*Get the Denominator and Numerator out of a Fraction*}
   194 text{*\noindent The selv written functions in e.g. \texttt{get\_denominator}
   195        should become a constant for the isabelle parser:*}
   196 
   197 consts
   198   get_denominator :: "real => real"
   199   get_numerator :: "real => real"
   200 
   201 text {*\noindent With the above definition we run into problems when we parse
   202         the Script \texttt{InverseZTransform}. This leads to \em ambiguous
   203         parse trees. \normalfont We avoid this by moving the definition
   204         to \ttfamily Rational.thy \normalfont and re-building Isac.
   205         \par \noindent ATTENTION: From now on \ttfamily 
   206         Build\_Inverse\_Z\_Transform \normalfont mimics a build from scratch;
   207         it only works due to re-building Isac several times (indicated 
   208         explicityl).
   209 *}
   210 
   211 ML {*
   212 (*
   213  *("get_denominator",
   214  *  ("Rational.get_denominator", eval_get_denominator ""))
   215  *)
   216 fun eval_get_denominator (thmid:string) _ 
   217 		      (t as Const ("Rational.get_denominator", _) $
   218               (Const ("Rings.inverse_class.divide", _) $ num $
   219                 denom)) thy = 
   220         SOME (mk_thmid thmid "" 
   221             (Print_Mode.setmp [] 
   222               (Syntax.string_of_term (thy2ctxt thy)) denom) "", 
   223 	          Trueprop $ (mk_equality (t, denom)))
   224   | eval_get_denominator _ _ _ _ = NONE; 
   225 *}
   226 text {* tests of eval_get_denominator see test/Knowledge/rational.sml*}
   227 
   228 text {*get numerator should also become a constant for the isabelle parser: *}
   229 
   230 ML {*
   231 fun eval_get_numerator (thmid:string) _ 
   232 		      (t as Const ("Rational.get_numerator", _) $
   233               (Const ("Rings.inverse_class.divide", _) $num
   234                 $denom )) thy = 
   235         SOME (mk_thmid thmid "" 
   236             (Print_Mode.setmp [] (Syntax.string_of_term (thy2ctxt thy)) num) "", 
   237 	          Trueprop $ (mk_equality (t, num)))
   238   | eval_get_numerator _ _ _ _ = NONE; 
   239 *}
   240 
   241 text {*
   242 We discovered severell problems by implementing the get_numerator function.
   243 Remember when putting new functions to Isac, put them in a thy file and rebuild 
   244 isac, also put them in the ruleset for the script!
   245 *}
   246 
   247 subsection {*solve equation*}
   248 text {*this type of equation if too general for the present program*}
   249 ML {*
   250 "----------- Minisubplb/100-init-rootp (*OK*)bl.sml ---------------------";
   251 val denominator = parseNEW ctxt "z^^^2 - 1/4*z - 1/8 = 0";
   252 val fmz = ["equality (z^^^2 - 1/4*z - 1/8 = (0::real))", "solveFor z","solutions L"];
   253 val (dI',pI',mI') =("Isac", ["univariate","equation"], ["no_met"]);
   254 (*                           ^^^^^^^^^^^^^^^^^^^^^^ TODO: ISAC determines type of eq*)
   255 *}
   256 text {*Does the Equation Match the Specification ?*}
   257 ML {*
   258 match_pbl fmz (get_pbt ["univariate","equation"]);
   259 *}
   260 ML {*Context.theory_name thy = "Isac"(*==================================================*)*}
   261 
   262 ML {*
   263 val denominator = parseNEW ctxt "-1 + -2 * z + 8 * z ^^^ 2 = 0";
   264 val fmz =                                            (*specification*)
   265   ["equality (-1 + -2 * z + 8 * z ^^^ 2 = (0::real))", (*equality*)
   266    "solveFor z",                                     (*bound variable*)
   267    "solutions L"];                                   (*identifier for solution*)
   268 
   269 val (dI',pI',mI') =
   270   ("Isac", ["abcFormula","degree_2","polynomial","univariate","equation"], ["no_met"]);
   271 *}
   272 text {*Does the Other Equation Match the Specification ?*}
   273 ML {*
   274 match_pbl fmz (get_pbt ["abcFormula","degree_2","polynomial","univariate","equation"]);
   275 *}
   276 text {*Solve Equation Stepwise*}
   277 ML {*
   278 *}
   279 ML {*
   280 val (p,_,f,nxt,_,pt) = CalcTreeTEST [(fmz, (dI',pI',mI'))];
   281 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   282 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   283 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   284 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   285 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   286 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   287 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   288 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   289 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   290 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
   291 val (p,_,f,nxt,_,pt) = me nxt p [] pt;         
   292 val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*nxt =..,Check_elementwise "Assumptions")*)
   293 val (p,_,f,nxt,_,pt) = me nxt p [] pt;         
   294 val (p,_,f,nxt,_,pt) = me nxt p [] pt; f2str f;
   295 (*[z = 1 / 2, z = -1 / 4]*)
   296 show_pt pt; 
   297 val SOME f = parseNEW ctxt "[z=1/2, z=-1/4]";
   298 *}
   299 
   300 subsection {*partial fraction decomposition*}
   301 subsubsection {*solution of the equation*}
   302 ML {*
   303 val SOME solutions = parseNEW ctxt "[z=1/2, z=-1/4]";
   304 term2str solutions;
   305 atomty solutions;
   306 *}
   307 
   308 subsubsection {*get solutions out of list*}
   309 text {*in isac's CTP-based programming language: let$ $s_1 = NTH 1$ solutions; $s_2 = NTH 2...$*}
   310 ML {*
   311 val Const ("List.list.Cons", _) $ s_1 $ (Const ("List.list.Cons", _) $
   312       s_2 $ Const ("List.list.Nil", _)) = solutions;
   313 term2str s_1;
   314 term2str s_2;
   315 *}
   316 
   317 ML {* (*Solutions as Denominator --> Denominator1 = z - Zeropoint1, Denominator2 = z-Zeropoint2,...*)
   318 val xx = HOLogic.dest_eq s_1;
   319 val s_1' = HOLogic.mk_binop "Groups.minus_class.minus" xx;
   320 val xx = HOLogic.dest_eq s_2;
   321 val s_2' = HOLogic.mk_binop "Groups.minus_class.minus" xx;
   322 term2str s_1';
   323 term2str s_2';
   324 *}
   325 text {* for the programming language a function 
   326   collecting all the above manipulations is helpful*}
   327 ML {*
   328 fun mk_minus_1 T = Free("-1", T); (*TODO DELETE WITH numbers_to_string*)
   329 fun flip_sign t = (*TODO improve for use in factors_from_solution: -(-1) etc*)
   330   let val minus_1 = t |> type_of |> mk_minus_1
   331   in HOLogic.mk_binop "Groups.times_class.times" (minus_1, t) end;
   332 fun fac_from_sol s =
   333   let val (lhs, rhs) = HOLogic.dest_eq s
   334   in HOLogic.mk_binop "Groups.minus_class.minus" (lhs, rhs) end;
   335 *}
   336 ML {*
   337 e_term
   338 *}
   339 ML {*
   340 fun mk_prod prod [] =
   341       if prod = e_term then error "mk_prod called with []" else prod
   342   | mk_prod prod (t :: []) =
   343       if prod = e_term then t else HOLogic.mk_binop "Groups.times_class.times" (prod, t)
   344   | mk_prod prod (t1 :: t2 :: ts) =
   345         if prod = e_term 
   346         then 
   347            let val p = HOLogic.mk_binop "Groups.times_class.times" (t1, t2)
   348            in mk_prod p ts end 
   349         else 
   350            let val p = HOLogic.mk_binop "Groups.times_class.times" (prod, t1)
   351            in mk_prod p (t2 :: ts) end 
   352 *}
   353 ML {*
   354 *}
   355 ML {*
   356 (*probably keept these test in test/Tools/isac/...
   357 (*mk_prod e_term [];*)
   358 
   359 val prod = mk_prod e_term [str2term "x + 123"]; 
   360 term2str prod = "x + 123";
   361 
   362 val sol = str2term "[z = 1 / 2, z = -1 / 4]";
   363 val sols = HOLogic.dest_list sol;
   364 val facs = map fac_from_sol sols;
   365 val prod = mk_prod e_term facs; 
   366 term2str prod = "(z + -1 * (1 / 2)) * (z + -1 * (-1 / 4))";
   367 
   368 val prod = mk_prod e_term [str2term "x + 1", str2term "x + 2", str2term "x + 3"]; 
   369 term2str prod = "(x + 1) * (x + 2) * (x + 3)";
   370 *)
   371 
   372 fun factors_from_solution sol = 
   373   let val ts = HOLogic.dest_list sol
   374   in mk_prod e_term (map fac_from_sol ts) end;
   375 (*
   376 val sol = str2term "[z = 1 / 2, z = -1 / 4]";
   377 val fs = factors_from_solution sol;
   378 term2str fs = "(z + -1 * (1 / 2)) * (z + -1 * (-1 / 4))"
   379 *)
   380 *}
   381 text {* This function needs to be packed such that it can be evaluated by the Lucas-Interpreter:
   382   # shift these functions into the related Equation.thy
   383   #  -- compare steps done with get_denominator above
   384   # done 02.12.2011 moved to PartialFractions.thy
   385   *}
   386 ML {*
   387 (*("factors_from_solution", ("Partial_Fractions.factors_from_solution", eval_factors_from_solution ""))*)
   388 fun eval_factors_from_solution (thmid:string) _
   389      (t as Const ("Partial_Fractions.factors_from_solution", _) $ sol) thy =
   390        ((let val prod = factors_from_solution sol
   391          in SOME (mk_thmid thmid ""
   392            (Print_Mode.setmp [] (Syntax.string_of_term (thy2ctxt thy)) prod) "",
   393                Trueprop $ (mk_equality (t, prod)))
   394          end)
   395        handle _ => NONE)
   396  | eval_factors_from_solution _ _ _ _ = NONE;
   397 *}
   398 
   399 text {*
   400 The tracing output of the calc tree after apllying this function was
   401 24 / factors_from_solution [z = 1/ 2, z = -1 / 4])]  and the next step
   402  val nxt = ("Empty_Tac", ...): tac'_).
   403 These observations indicate, that the Lucas-Interpreter (LIP) does 
   404 not know how to evaluate factors_from_solution, so there is something 
   405 wrong or missing.
   406 
   407 # First we isolate the difficulty in the program as follows:
   408   :
   409   "      (L_L::bool list) = (SubProblem (PolyEq'," ^
   410   "          [abcFormula,degree_2,polynomial,univariate,equation],[no_met])" ^
   411   "        [BOOL equ, REAL zzz]);              " ^
   412   "      (facs::real) = factors_from_solution L_L;" ^
   413   "      (foo::real) = Take facs" ^
   414   :
   415 and see
   416   [
   417   (([], Frm), Problem (Isac, [inverse, Z_Transform, SignalProcessing])),
   418   (([1], Frm), X z = 3 / (z - 1 / 4 + -1 / 8 * (1 / z))),
   419   (([1], Res), ?X' z = 3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z)))),
   420   (([2], Res), ?X' z = 24 / (-1 + -2 * z + 8 * z ^^^ 2)),
   421   (([3], Pbl), solve (-1 + -2 * z + 8 * z ^^^ 2 = 0, z)),
   422   (([3,1], Frm), -1 + -2 * z + 8 * z ^^^ 2 = 0),
   423   (([3,1], Res), z = (- -2 + sqrt (-2 ^^^ 2 - 4 * 8 * -1)) / (2 * 8) |
   424                  z = (- -2 - sqrt (-2 ^^^ 2 - 4 * 8 * -1)) / (2 * 8)),
   425   (([3,2], Res), z = 1 / 2 | z = -1 / 4),
   426   (([3,3], Res), [z = 1 / 2, z = -1 / 4]),
   427   (([3,4], Res), [z = 1 / 2, z = -1 / 4]),
   428   (([3], Res), [z = 1 / 2, z = -1 / 4]),
   429   (([4], Frm), factors_from_solution [z = 1 / 2, z = -1 / 4])]
   430 in particular that
   431   (([3], Pbl), solve (-1 + -2 * z + 8 * z ^^^ 2 = 0, z)),
   432 shows the equation which has been created in the program by
   433   "      (denom::real) = get_denominator funterm;" ^ (*get_denominator*)
   434   "      (equ::bool) = (denom = (0::real));" ^
   435 # 'get_denominator' has been evaluated successfully, but not factors_from_solution.
   436 So we stepwise compare with an analogous case, get_denominator 
   437 successfully done above: We know that LIP evaluates expressions in the 
   438 program by use of the "srls", so we
   439 # try to get the original srls
   440 
   441   val {srls, ...} = get_met ["SignalProcessing","Z_Transform","inverse"];
   442 
   443 # create 2 good example terms
   444   val SOME t1 = parseNEW ctxt "get_denominator ((111::real) / 222)";
   445   val SOME t2 = parseNEW ctxt "factors_from_solution [(z::real) = 1 / 2, z = -1 / 4]";
   446 
   447 # rewrite the terms using srls
   448   rewrite_set_ thy true srls t1;
   449   rewrite_set_ thy true srls t2;
   450 
   451 and we see a difference: t1 gives SOME, t2 gives NONE.
   452 Now we look at the srls:
   453   val srls = Rls {id="srls_InverseZTransform",
   454     :
   455     rules =
   456       [
   457         :
   458         Calc("Rational.get_numerator",
   459           eval_get_numerator "Rational.get_numerator"),
   460         Calc("Partial_Fractions.factors_from_solution",
   461           eval_factors_from_solution "Partial_Fractions.factors_from_solution")
   462       ],
   463     :
   464 
   465 Here everthing is perfect. So the error can only be in the SML code of eval_factors_from_solution.
   466 We try to check the code with an existing test; since the code is in
   467 
   468   src/Tools/isac/Knowledge/Partial_Fractions.thy
   469 
   470 the test should be in
   471 
   472   test/Tools/isac/Knowledge/partial_fractions.sml
   473 
   474 -------------------------------------------------------------------------------
   475 After updating the function get_factors_from solution to a new version and 
   476 putting a testcase to Partial_Fractions.sml we tried again to evaluate the
   477 term with the same result.
   478 We opened the test Test_Isac.thy and saw that everything is working fine.
   479 Also we checked that the test partial_fractions.sml is used in Test_Isac.thy 
   480 
   481 -->  use "Knowledge/partial_fractions.sml"
   482 
   483 and Partial_Fractions.thy is part is part of isac by evaluating
   484 
   485 val thy = @{theory Isac};
   486 
   487 after rebuilding isac again it worked
   488 
   489 *}
   490 
   491 subsubsection {*build expression*}
   492 text {*in isac's CTP-based programming language: let s_1 = Take numerator / (s_1 * s_2)*}
   493 ML {*
   494 (*The Main Denominator is the multiplikation of the partial fraction denominators*)
   495 val denominator' = HOLogic.mk_binop "Groups.times_class.times" (s_1', s_2') ;
   496 val SOME numerator = parseNEW ctxt "3::real";
   497 
   498 val expr' = HOLogic.mk_binop "Rings.inverse_class.divide" (numerator, denominator');
   499 term2str expr';
   500 *}
   501 
   502 subsubsection {*Ansatz - create partial fractions out of our expression*}
   503 ML {*Context.theory_name thy = "Isac"*}
   504 
   505 axiomatization where
   506   ansatz2: "n / (a*b) = A/a + B/(b::real)" and
   507   multiply_eq2: "((n::real) / (a*b) = A/a + B/b) = (a*b*(n  / (a*b)) = a*b*(A/a + B/b::real))"  
   508 
   509 ML {*
   510 (*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*)
   511 val SOME (t1,_) = rewrite_ @{theory Isac} e_rew_ord e_rls false @{thm ansatz2} expr';
   512 term2str t1; atomty t1;
   513 val eq1 = HOLogic.mk_eq (expr', t1);
   514 term2str eq1;
   515 *}
   516 ML {*
   517 (*eliminate the demoninators by multiplying the left and the right side with the main denominator*)
   518 val SOME (eq2,_) = rewrite_ @{theory Isac} e_rew_ord e_rls false @{thm multiply_eq2} eq1;
   519 term2str eq2;
   520 *}
   521 ML {*
   522 (*simplificatoin*)
   523 val SOME (eq3,_) = rewrite_set_ @{theory Isac} false norm_Rational eq2;
   524 term2str eq3; (*?A ?B not simplified*)
   525 *}
   526 ML {*
   527 val SOME fract1 =
   528   parseNEW ctxt "(z - 1 / 2) * (z - -1 / 4) * (A / (z - 1 / 2) + B / (z - -1 / 4))"; (*A B !*)
   529 val SOME (fract2,_) = rewrite_set_ @{theory Isac} false norm_Rational fract1;
   530 term2str fract2 = "(A + -2 * B + 4 * A * z + 4 * B * z) / 4";
   531 (*term2str fract2 = "A * (1 / 4 + z) + B * (-1 / 2 + z)" would be more traditional*)
   532 *}
   533 ML {*
   534 val (numerator, denominator) = HOLogic.dest_eq eq3;
   535 val eq3' = HOLogic.mk_eq (numerator, fract1); (*A B !*)
   536 term2str eq3';
   537 (*MANDATORY: simplify (and remove denominator) otherwise 3 = 0*)
   538 val SOME (eq3'' ,_) = rewrite_set_ @{theory Isac} false norm_Rational eq3';
   539 term2str eq3'';
   540 *}
   541 ML {*Context.theory_name thy = "Isac"(*==================================================*)*}
   542 
   543 subsubsection {*Build a rule-set for ansatz*}
   544 text {* the "ansatz" rules violate the principle that each variable on 
   545   the right-hand-side must also occur on the left-hand-side of the rule:
   546   A, B, etc don't.
   547   Thus the rewriter marks these variables with question marks: ?A, ?B, etc.
   548   These question marks can be dropped by "fun drop_questionmarks".
   549   *}
   550 ML {*
   551 val ansatz_rls = prep_rls(
   552   Rls {id = "ansatz_rls", preconds = [], rew_ord = ("dummy_ord",dummy_ord), 
   553 	  erls = e_rls, srls = Erls, calc = [],
   554 	  rules = 
   555 	   [Thm ("ansatz2",num_str @{thm ansatz2}),
   556 	    Thm ("multiply_eq2",num_str @{thm multiply_eq2})
   557 	   ], 
   558 	 scr = EmptyScr});
   559 *}
   560 ML {*
   561 val SOME (ttttt,_) = rewrite_set_ @{theory Isac} false ansatz_rls expr';
   562 *}
   563 ML {*
   564 term2str expr' = "3 / ((z - 1 / 2) * (z - -1 / 4))";
   565 term2str ttttt = "?A / (z - 1 / 2) + ?B / (z - -1 / 4)";
   566 *}
   567 
   568 
   569 subsubsection {*get first koeffizient*}
   570 
   571 ML {*
   572 (*substitude z with the first zeropoint to get A*)
   573 val SOME (eq4_1,_) = rewrite_terms_ @{theory Isac} e_rew_ord e_rls [s_1] eq3'';
   574 term2str eq4_1;
   575 
   576 val SOME (eq4_2,_) = rewrite_set_ @{theory Isac} false norm_Rational eq4_1;
   577 term2str eq4_2;
   578 
   579 val fmz = ["equality (3 = 3 * A / (4::real))", "solveFor A","solutions L"];
   580 val (dI',pI',mI') =("Isac", ["univariate","equation"], ["no_met"]);
   581 (*solve the simple linear equilation for A TODO: return eq, not list of eq*)
   582 val (p,_,fa,nxt,_,pt) = CalcTreeTEST [(fmz, (dI',pI',mI'))];
   583 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Add_Given "equality (3 = 3 * A / 4)"*)
   584 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (* Add_Given "solveFor A"*)
   585 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Add_Find "solutions L"*)
   586 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Specify_Theory "Isac"*)
   587 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Specify_Problem ["normalize", "polynomial", 
   588                                           "univariate", "equation"])*)
   589 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (* Specify_Method ["PolyEq", "normalize_poly"]*)
   590 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Apply_Method ["PolyEq", "normalize_poly"]*)
   591 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Rewrite ("all_left", "PolyEq.all_left")*)
   592 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Rewrite_Set_Inst (["(bdv, A)"], "make_ratpoly_in")*)
   593 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Rewrite_Set "polyeq_simplify"*)
   594 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (**)
   595 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (**)
   596 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Add_Given "equality (3 + -3 / 4 * A = 0)"*)
   597 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Add_Given "solveFor A"*)
   598 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Add_Find "solutions A_i"*)
   599 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (**)
   600 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (**)
   601 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (**)
   602 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Apply_Method ["PolyEq", "solve_d1_polyeq_equation"]*)
   603 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Rewrite_Set_Inst (["(bdv, A)"], "d1_polyeq_simplify")*)
   604 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Rewrite_Set "polyeq_simplify"*)
   605 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Rewrite_Set "norm_Rational_parenthesized"*)
   606 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Or_to_List*)
   607 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Check_elementwise "Assumptions"*)
   608 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Check_Postcond ["degree_1", "polynomial", 
   609                                           "univariate", "equation"]*)
   610 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*Check_Postcond ["normalize", "polynomial", 
   611                                           "univariate", "equation"]*)
   612 val (p,_,fa,nxt,_,pt) = me nxt p [] pt; (*End_Proof'*)
   613 f2str fa;
   614 *}
   615 
   616 subsubsection {*get second koeffizient*}
   617 ML {*thy*}
   618 
   619 ML {*
   620 (*substitude z with the second zeropoint to get B*)
   621 val SOME (eq4b_1,_) = rewrite_terms_ @{theory Isac} e_rew_ord e_rls [s_2] eq3'';
   622 term2str eq4b_1;
   623 
   624 val SOME (eq4b_2,_) = rewrite_set_ @{theory Isac} false norm_Rational eq4b_1;
   625 term2str eq4b_2;
   626 *}
   627 ML {*
   628 (*solve the simple linear equilation for B TODO: return eq, not list of eq*)
   629 val fmz = ["equality (3 = -3 * B / (4::real))", "solveFor B","solutions L"];
   630 val (dI',pI',mI') =("Isac", ["univariate","equation"], ["no_met"]);
   631 val (p,_,fb,nxt,_,pt) = CalcTreeTEST [(fmz, (dI',pI',mI'))];
   632 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   633 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   634 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   635 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   636 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   637 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   638 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   639 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   640 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   641 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   642 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   643 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   644 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   645 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   646 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   647 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   648 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   649 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   650 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   651 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   652 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   653 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   654 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   655 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   656 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   657 val (p,_,fb,nxt,_,pt) = me nxt p [] pt;
   658 val (p,_,fb,nxt,_,pt) = me nxt p [] pt; 
   659 f2str fb;
   660 *}
   661 
   662 ML {* (*check koeffizients*)
   663 if f2str fa = "[A = 4]" then () else error "part.fract. eq4_1";
   664 if f2str fb = "[B = -4]" then () else error "part.fract. eq4_1";
   665 *}
   666 
   667 subsubsection {*substitute expression with solutions*}
   668 ML {*
   669 *}
   670 ML {*thy*}
   671 
   672 section {*Implement the Specification and the Method \label{spec-meth}*}
   673 text{*==============================================*}
   674 subsection{*Define the Field Descriptions for the specification*}
   675 consts
   676   filterExpression  :: "bool => una"
   677   stepResponse      :: "bool => una"
   678 
   679 subsection{*Define the Specification*}
   680 ML {*
   681 store_pbt
   682  (prep_pbt thy "pbl_SP" [] e_pblID
   683  (["SignalProcessing"], [], e_rls, NONE, []));
   684 store_pbt
   685  (prep_pbt thy "pbl_SP_Ztrans" [] e_pblID
   686  (["Z_Transform","SignalProcessing"], [], e_rls, NONE, []));
   687 *}
   688 ML {*thy*}
   689 ML {*
   690 store_pbt
   691  (prep_pbt thy "pbl_SP_Ztrans_inv" [] e_pblID
   692  (["inverse", "Z_Transform", "SignalProcessing"],
   693   [("#Given" ,["filterExpression X_eq"]),
   694    ("#Find"  ,["stepResponse n_eq"])
   695   ],
   696   append_rls "e_rls" e_rls [(*for preds in where_*)], NONE, 
   697   [["SignalProcessing","Z_Transform","inverse"]]));
   698 
   699 show_ptyps();
   700 get_pbt ["inverse","Z_Transform","SignalProcessing"];
   701 *}
   702 
   703 subsection {*Define Name and Signature for the Method*}
   704 consts
   705   InverseZTransform :: "[bool, bool] => bool"
   706     ("((Script InverseZTransform (_ =))// (_))" 9)
   707 
   708 subsection {*Setup Parent Nodes in Hierarchy of Method*}
   709 ML {*
   710 store_met
   711  (prep_met thy "met_SP" [] e_metID
   712  (["SignalProcessing"], [],
   713    {rew_ord'="tless_true", rls'= e_rls, calc = [], srls = e_rls, prls = e_rls,
   714     crls = e_rls, nrls = e_rls}, "empty_script"));
   715 store_met
   716  (prep_met thy "met_SP_Ztrans" [] e_metID
   717  (["SignalProcessing", "Z_Transform"], [],
   718    {rew_ord'="tless_true", rls'= e_rls, calc = [], srls = e_rls, prls = e_rls,
   719     crls = e_rls, nrls = e_rls}, "empty_script"));
   720 *}
   721 ML {*
   722 store_met
   723  (prep_met thy "met_SP_Ztrans_inv" [] e_metID
   724  (["SignalProcessing", "Z_Transform", "inverse"], 
   725   [("#Given" ,["filterExpression X_eq"]),
   726    ("#Find"  ,["stepResponse n_eq"])
   727   ],
   728    {rew_ord'="tless_true", rls'= e_rls, calc = [], srls = e_rls, prls = e_rls,
   729     crls = e_rls, nrls = e_rls},
   730   "empty_script"
   731  ));
   732 *}
   733 ML {*
   734 store_met
   735  (prep_met thy "met_SP_Ztrans_inv" [] e_metID
   736  (["SignalProcessing", "Z_Transform", "inverse"], 
   737   [("#Given" ,["filterExpression X_eq"]),
   738    ("#Find"  ,["stepResponse n_eq"])
   739   ],
   740    {rew_ord'="tless_true", rls'= e_rls, calc = [], srls = e_rls, prls = e_rls,
   741     crls = e_rls, nrls = e_rls},
   742   "Script InverseZTransform (Xeq::bool) =" ^
   743   " (let X = Take Xeq;" ^
   744   "      X = Rewrite ruleZY False X" ^
   745   "  in X)"
   746  ));
   747 *}
   748 ML {*
   749 show_mets();
   750 *}
   751 ML {*
   752 get_met ["SignalProcessing","Z_Transform","inverse"];
   753 *}
   754 
   755 section {*Program in CTP-based language \label{prog-steps}*}
   756 text{*=================================*}
   757 subsection {*Stepwise extend Program*}
   758 ML {*
   759 val str = 
   760 "Script InverseZTransform (Xeq::bool) =" ^
   761 " Xeq";
   762 *}
   763 ML {*
   764 val str = 
   765 "Script InverseZTransform (Xeq::bool) =" ^ (*(1/z) instead of z ^^^ -1*)
   766 " (let X = Take Xeq;" ^
   767 "      X' = Rewrite ruleZY False X;" ^ (*z * denominator*)
   768 "      X' = (Rewrite_Set norm_Rational False) X'" ^ (*simplify*)
   769 "  in X)";
   770 (*NONE*)
   771 "Script InverseZTransform (Xeq::bool) =" ^ (*(1/z) instead of z ^^^ -1*)
   772 " (let X = Take Xeq;" ^
   773 "      X' = Rewrite ruleZY False X;" ^ (*z * denominator*)
   774 "      X' = (Rewrite_Set norm_Rational False) X';" ^ (*simplify*)
   775 "      X' = (SubProblem (Isac',[pqFormula,degree_2,polynomial,univariate,equation], [no_met])   " ^
   776     "                 [BOOL e_e, REAL v_v])" ^
   777 "  in X)";
   778 *}
   779 ML {*
   780 val str = 
   781 "Script InverseZTransform (Xeq::bool) =" ^ (*(1/z) instead of z ^^^ -1*)
   782 " (let X = Take Xeq;" ^
   783 "      X' = Rewrite ruleZY False X;" ^ (*z * denominator*)
   784 "      X' = (Rewrite_Set norm_Rational False) X';" ^ (*simplify*)
   785 "      funterm = rhs X'" ^ (*drop X'= for equation solving*)
   786 "  in X)";
   787 *}
   788 ML {*
   789 val str = 
   790 "Script InverseZTransform (X_eq::bool) =" ^ (*(1/z) instead of z ^^^ -1*)
   791 " (let X = Take X_eq;" ^
   792 "      X' = Rewrite ruleZY False X;" ^ (*z * denominator*)
   793 "      X' = (Rewrite_Set norm_Rational False) X';" ^ (*simplify*)
   794 "      (X'_z::real) = lhs X';" ^
   795 "      (z::real) = argument_in X'_z;" ^
   796 "      (funterm::real) = rhs X';" ^ (*drop X' z = for equation solving*)
   797 "      (denom::real) = get_denominator funterm;" ^ (*get_denominator*)
   798 "      (equ::bool) = (denom = (0::real));" ^
   799 "      (L_L::bool list) =                                    " ^
   800 "            (SubProblem (Test',                            " ^
   801 "                         [linear,univariate,equation,test]," ^
   802 "                         [Test,solve_linear])              " ^
   803 "                        [BOOL equ, REAL z])              " ^
   804 "  in X)"
   805 ;
   806 
   807 parse thy str;
   808 val sc = ((inst_abs thy) o term_of o the o (parse thy)) str;
   809 atomty sc;
   810 
   811 *}
   812 
   813 text {*
   814 This ruleset contains all functions that are in the script; 
   815 The evaluation of the functions is done by rewriting using this ruleset.
   816 *}
   817 
   818 ML {*
   819 val srls = Rls {id="srls_InverseZTransform", 
   820 		  preconds = [], rew_ord = ("termlessI",termlessI), 
   821 		  erls = append_rls "erls_in_srls_InverseZTransform" e_rls
   822 				    [(*for asm in NTH_CONS ...*) Calc ("Orderings.ord_class.less",eval_equ "#less_"),
   823 				     (*2nd NTH_CONS pushes n+-1 into asms*) Calc("Groups.plus_class.plus", eval_binop "#add_")
   824 				    ], 
   825   srls = Erls, calc = [],
   826 		  rules =
   827     [Thm ("NTH_CONS",num_str @{thm NTH_CONS}),
   828 			     Calc("Groups.plus_class.plus", eval_binop "#add_"),
   829 			     Thm ("NTH_NIL",num_str @{thm NTH_NIL}),
   830 			     Calc("Tools.lhs", eval_lhs"eval_lhs_"), (*<=== ONLY USED*)
   831 			     Calc("Tools.rhs", eval_rhs"eval_rhs_"), (*<=== ONLY USED*)
   832 			     Calc("Atools.argument'_in", eval_argument_in "Atools.argument'_in"),
   833      Calc("Rational.get_denominator", eval_get_denominator "#get_denominator"),
   834      Calc("Rational.get_numerator", eval_get_numerator "#get_numerator"),
   835      Calc("Partial_Fractions.factors_from_solution",
   836        eval_factors_from_solution "#factors_from_solution"),
   837      Calc("Partial_Fractions.drop_questionmarks", eval_drop_questionmarks "#drop_?")
   838 			    ],
   839 		  scr = EmptyScr};
   840 *}
   841 
   842 
   843 subsection {*Store Final Version of Program for Execution*}
   844 
   845 ML {*
   846 store_met
   847  (prep_met thy "met_SP_Ztrans_inv" [] e_metID
   848  (["SignalProcessing", "Z_Transform", "inverse"], 
   849   [("#Given" ,["filterExpression X_eq"]),
   850    ("#Find"  ,["stepResponse n_eq"])
   851   ],
   852    {rew_ord'="tless_true", rls'= e_rls, calc = [], srls = srls, 
   853     prls = e_rls,
   854     crls = e_rls, nrls = e_rls},
   855       "Script InverseZTransform (X_eq::bool) =" ^ (*(1/z) instead of z ^^^ -1*)
   856       " (let X = Take X_eq;" ^
   857 (*([1], Frm), X z = 3 / (z - 1 / 4 + -1 / 8 * (1 / z))*)
   858       "      X' = Rewrite ruleZY False X;" ^ (*z * denominator*)
   859 (*([1], Res), ?X' z = 3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z)))*)
   860       "      (num_orig::real) = get_numerator (rhs X');"^
   861       "      X' = (Rewrite_Set norm_Rational False) X';" ^ (*simplify*)
   862 (*([2], Res), ?X' z = 24 / (-1 + -2 * z + 8 * z ^^^ 2)*)
   863       "      (X'_z::real) = lhs X';" ^ (**)
   864       "      (zzz::real) = argument_in X'_z;" ^ (**)
   865       "      (funterm::real) = rhs X';" ^ (*drop X' z = for equation solving*)
   866       "      (denom::real) = get_denominator funterm;" ^ (*get_denominator*)
   867       "      (num::real) = get_numerator funterm; " ^ (*get_numerator*)
   868       "      (equ::bool) = (denom = (0::real));" ^
   869       "      (L_L::bool list) = (SubProblem (PolyEq'," ^
   870       "         [abcFormula,degree_2,polynomial,univariate,equation],[no_met])" ^
   871       "         [BOOL equ, REAL zzz]);              " ^
   872 
   873 (*([3], Pbl), solve (-1 + -2 * z + 8 * z ^^^ 2 = 0, z)*)
   874 (*([3], Res), [z = 1 / 2, z = -1 / 4]*)
   875 
   876       "      (facs::real) = factors_from_solution L_L;" ^
   877       "      (eql::real) = Take (num_orig / facs);" ^                                 (*([4], Frm), 24 / ((z + -1 * (1 / 2)) * (z + -1 * (-1 / 4)))*)
   878 
   879       "      (eqr::real) = (Try (Rewrite_Set ansatz_rls False)) eql;"^           (*([4], Res), ?A / (z + -1 * (1 / 2)) + ?B / (z + -1 * (-1 / 4))*)
   880 
   881       "      (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))*)
   882 
   883       "      eq = (Try (Rewrite_Set equival_trans False)) eq;"^                  (*([5], Res), 24 = ?A * (z + -1 * (-1 / 4)) + ?B * (z + -1 * (1 / 2))*)
   884 
   885       "      eq = drop_questionmarks eq;"^
   886       "      (z1::real) = (rhs (NTH 1 L_L));"^                                   (*prepare equliation for a - eq_a therfor subsitude z with solution 1 - z1*)
   887       "      (z2::real) = (rhs (NTH 2 L_L));"^
   888  
   889       "      (eq_a::bool) = Take eq;"^
   890       "      eq_a = (Substitute [zzz=z1]) eq;"^                          (*([6], Res), 24 = ?A * (1 / 2 + -1 * (-1 / 4)) + ?B * (1 / 2 + -1 * (1 / 2))*)
   891       "      eq_a = (Rewrite_Set norm_Rational False) eq_a;"^                    (*([7], Res), 24 = ?A * 3 / 4*)
   892       "      (sol_a::bool list) = (SubProblem (Isac'," ^
   893       "          [univariate,equation],[no_met])" ^
   894       "        [BOOL eq_a, REAL (A::real)]);"^
   895       "      (a::real) = (rhs(NTH 1 sol_a));"^
   896 
   897       "      (eq_b::bool) = Take eq;"^
   898       "      eq_b =  (Substitute [zzz=z2]) eq_b;"^
   899       "      eq_b = (Rewrite_Set norm_Rational False) eq_b;"^
   900       "      (sol_b::bool list) = (SubProblem (Isac'," ^
   901       "          [univariate,equation],[no_met])" ^
   902       "        [BOOL eq_b, REAL (B::real)]);"^
   903       "      (b::real) = (rhs(NTH 1 sol_b));"^
   904 
   905 
   906       "      eqr = drop_questionmarks eqr;"^
   907       "      (pbz::real) = Take eqr;"^
   908       "      pbz = ((Substitute [A=a]) pbz);"^
   909       "      pbz = ((Substitute [B=b]) pbz);"^
   910 
   911       "      pbz = Rewrite ruleYZ False pbz;"^
   912       "      pbz = drop_questionmarks pbz;"^
   913 
   914       "      (iztrans::real) = Take pbz;"^
   915       "      iztrans = (Rewrite_Set inverse_z False) iztrans;"^
   916       "      iztrans = drop_questionmarks iztrans;"^
   917       "      (n_eq::bool) = Take (X_n = iztrans)"^ 
   918 
   919       "  in n_eq)" 
   920       ));
   921 *}
   922 
   923 
   924 subsection {*Check the Program*}
   925 
   926 subsubsection {*Check the formalization*}
   927 ML {*
   928 val fmz = ["filterExpression (X  = 3 / (z - 1/4 + -1/8 * (1/(z::real))))", 
   929   "stepResponse (x[n::real]::bool)"];
   930 val (dI,pI,mI) = ("Isac", ["inverse", "Z_Transform", "SignalProcessing"], 
   931   ["SignalProcessing","Z_Transform","inverse"]);
   932 
   933 val ([(1, [1], "#Given", Const ("Inverse_Z_Transform.filterExpression", _),
   934             [Const ("HOL.eq", _) $ _ $ _]),
   935            (2, [1], "#Find", Const ("Inverse_Z_Transform.stepResponse", _),
   936             [Free ("x", _) $ _])],
   937           _) = prep_ori fmz thy ((#ppc o get_pbt) pI);
   938 *}
   939 ML {*
   940 val Script sc = (#scr o get_met) ["SignalProcessing","Z_Transform","inverse"];
   941 atomty sc;
   942 *}
   943 
   944 subsubsection {*Stepwise check the program\label{sec:stepcheck}*}
   945 ML {*
   946 trace_rewrite := false;
   947 trace_script := false; print_depth 9;
   948 val fmz = ["filterExpression (X z = 3 / (z - 1/4 + -1/8 * (1/(z::real))))", 
   949   "stepResponse (x[n::real]::bool)"];
   950 val (dI,pI,mI) = ("Isac", ["inverse", "Z_Transform", "SignalProcessing"], 
   951   ["SignalProcessing","Z_Transform","inverse"]);
   952 val (p,_,f,nxt,_,pt)  = CalcTreeTEST [(fmz, (dI,pI,mI))];
   953 (*([], Frm), Problem (Isac, [inverse, Z_Transform, SignalProcessing])*)
   954 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "Add_Given";
   955 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "Add_Find";
   956 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "Specify_Theory";
   957 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "Specify_Problem";
   958 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "Specify_Method";
   959 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "nxt = Apply_Method";
   960 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!*)
   961 (*([1], Frm), X z = 3 / (z - 1 / 4 + -1 / 8 * (1 / z))*)
   962 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)))";
   963 (*([1], Res), ?X' z = 3 / (z * (z - 1 / 4 + -1 / 8 * (1 / z)))*)
   964 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "nxt = SubProblem";
   965 (*([2], Res), ?X' z = 24 / (-1 + -2 * z + 8 * z ^^^ 2)*)
   966 *}
   967 text {* Instead of nxt = Subproblem above we had Empty_Tac; the search for the reason 
   968   considered the following points:
   969   # what shows show_pt pt; ...
   970     (([2], Res), ?X' z = 24 / (-1 + -2 * z + 8 * z ^^^ 2))] ..calculation ok,
   971     but no "next" step found: should be "nxt = Subproblem" ?!?
   972   # what shows trace_script := true; we read bottom up ...
   973     @@@ next   leaf 'SubProbfrom
   974      (PolyEq', [abcFormula, degree_2, polynomial, univariate, equation],
   975       no_meth)
   976      [BOOL equ, REAL z]' ---> STac 'SubProblem
   977      (PolyEq', [abcFormula, degree_2, polynomial, univariate, equation],
   978       no_meth)
   979      [BOOL (-1 + -2 * z + 8 * z ^^^ 2 = 0), REAL z]'
   980     ... and see the SubProblem with correct arguments from searching next step
   981     (program text !!!--->!!! STac (script tactic) with arguments evaluated.)
   982   # do we have the right Script ...difference in the argumentsdifference in the arguments
   983     val Script s = (#scr o get_met) ["SignalProcessing","Z_Transform","inverse"];
   984     writeln (term2str s);
   985     ... shows the right script.difference in the arguments
   986   # test --- why helpless here ? --- shows: replace no_meth by [no_meth] in Script
   987 *}
   988 
   989 ML {*
   990 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "nxt = Model_Problem";
   991 (*([3], Pbl), solve (-1 + -2 * z + 8 * z ^^^ 2 = 0, z)*)
   992 *}
   993 text {* Instead of nxt = Model_Problem above we had Empty_Tac; the search for the reason 
   994   considered the following points:difference in the arguments
   995   # comparison with subsection { *solve equation* }: there solving this equation works,
   996     so there must be some difference in the arguments of the Subproblem:
   997     RIGHT: we had [no_meth] instead of [no_met] ;-))
   998 *}
   999 ML {*
  1000 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "nxt = Add_Given equality (-1 + -2 * z + 8 * z ^^^ 2 = 0)";
  1001 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "nxt = Add_Given solveFor z";
  1002 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "nxt = Add_Find solutions z_i";
  1003 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "nxt = Specify_Theory Isac";
  1004 *}
  1005 
  1006 text {* We had "nxt = Empty_Tac instead Specify_Theory; 
  1007   the search for the reason considered the following points:
  1008   # was there an error message ? NO --ok
  1009   # has "nxt = Add_Find" been inserted in pt: get_obj g_pbl pt (fst p); YES --ok
  1010   # what is the returned "formula": print_depth 999; f; print_depth 999; --
  1011     {Find = [Correct "solutions z_i"], With = [], 
  1012      Given = [Correct "equality (-1 + -2 * z + 8 * z ^^^ 2 = 0)", Correct "solveFor z"],
  1013      Where = [False "matches (z = 0) (-1 + -2 * z + 8 * z ^^^ 2 = 0) |\n
  1014                      matches (?b * z = 0) (-1 + -2 * z + 8 * z ^^^ 2 = 0) |\n
  1015                      matches (?a + z = 0) (-1 + -2 * z + 8 * z ^^^ 2 = 0) |\n
  1016                      matches (?a + ?b * z = 0) (-1 + -2 * z + 8 * z ^^^ 2 = 0)"],
  1017      Relate = []}
  1018      -- the only False is the reason: the Where (the precondition) is False for good reasons:
  1019      the precondition seems to check for linear equations, not for the one we want to solve!
  1020   Removed this error by correcting the Script
  1021   from SubProblem (PolyEq', [linear,univariate,equation,test], [Test,solve_linear]
  1022   to   SubProblem (PolyEq', [abcFormula,degree_2,polynomial,univariate,equation],
  1023                    [PolyEq,solve_d2_polyeq_abc_equation]
  1024   You find the appropriate type of equation at
  1025     http://www.ist.tugraz.at/projects/isac/www/kbase/pbl/index_pbl.html
  1026   and the respective method in Knowledge/PolyEq.thy at the respective store_pbt.
  1027   Or you leave the selection of the appropriate type to isac as done in the final Script ;-))
  1028 *}
  1029 ML {*
  1030 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "nxt = Specify_Problem [abcFormula,...";
  1031 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "nxt = Specify_Method [PolyEq,solve_d2_polyeq_abc_equation";
  1032 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "nxt = Apply_Method [PolyEq,solve_d2_polyeq_abc_equation";
  1033 val (p,_,f,nxt,_,pt) = me nxt p [] pt; "nxt = Rewrite_Set_Inst ([(bdv, z)], d2_polyeq_abcFormula_simplify";
  1034 (*([3,1], Frm), -1 + -2 * z + 8 * z ^^^ 2 = 0)*)
  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 (*([3,4], Res), [z = 1 / 2, z = -1 / 4])*)
  1040 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1041 (*([3], Res), [z = 1 / 2, z = -1 / 4]*)
  1042 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1043 (*([4], Frm), 24 / ((z + -1 * (1 / 2)) * (z + -1 * (-1 / 4)))*)
  1044 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1045 (*([4], Res), ?A / (z + -1 * (1 / 2)) + ?B / (z + -1 * (-1 / 4))*)
  1046 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1047 (*([5], Frm), 24 / ((z + -1 * (1 / 2)) * (z + -1 * (-1 / 4))) = ?A / (z + -1 * (1 / 2)) + ?B / (z + -1 * (-1 / 4))*)
  1048 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1049 (*([5], Res), 24 = ?A * (z + -1 * (-1 / 4)) + ?B * (z + -1 * (1 / 2))*)
  1050 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1051 (*([6], Res), 24 = A * (1 / 2 + -1 * (-1 / 4)) + B * (1 / 2 + -1 * (1 / 2))*)
  1052 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1053 (*([7], Res), 24 = A * 3 / 4*)
  1054 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1055 (*([8], Pbl), solve (24 = 3 * A / 4, A)*)
  1056 val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*Add_Given "equality (24 = 3 * A / 4)"*)
  1057 val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*Add_Given "solveFor A"*)
  1058 val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*Add_Find "solutions A_i"*)
  1059 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", 
  1060                                          "univariate", "equation"]*)
  1061 val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*Specify_Method ["PolyEq", "normalize_poly"]*)
  1062 val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*Apply_Method ["PolyEq", "normalize_poly"]*)
  1063 val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*Rewrite ("all_left", "PolyEq.all_left")*)
  1064 val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*Rewrite_Set_Inst (["(bdv, A)"], "make_ratpoly_in")*)
  1065 val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*Rewrite_Set "polyeq_simplify"*)
  1066 val (p,_,f,nxt,_,pt) = me nxt p [] pt; (*Subproblem ("Isac", ["degree_1", "polynomial", 
  1067                                          "univariate", "equation"])*)
  1068 *}
  1069 
  1070 
  1071 ML {*
  1072 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1073 show_pt pt;
  1074 *}
  1075 
  1076 ML {*
  1077 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1078 show_pt pt;
  1079 *}
  1080 
  1081 ML {*
  1082 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1083 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1084 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1085 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1086 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1087 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1088 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1089 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1090 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1091 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1092 *}
  1093 
  1094 ML {*
  1095 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1096 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1097 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1098 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1099 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1100 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1101 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1102 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1103 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1104 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1105 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1106 *}
  1107 
  1108 ML {*
  1109 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1110 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1111 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1112 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1113 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1114 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1115 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1116 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1117 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1118 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1119 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1120 *}
  1121 
  1122 ML {*
  1123 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1124 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1125 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1126 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1127 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1128 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1129 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1130 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1131 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1132 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1133 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1134 *}
  1135 
  1136 ML {*
  1137 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1138 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1139 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1140 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1141 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1142 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1143 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1144 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1145 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1146 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1147 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1148 *}
  1149 
  1150 ML {*
  1151 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1152 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1153 *}
  1154 
  1155 ML {*
  1156 trace_script := true;
  1157 val (p,_,f,nxt,_,pt) = me nxt p [] pt;
  1158 show_pt pt;
  1159 *}
  1160 
  1161 
  1162 section {*Write Tests for Crucial Details*}
  1163 text{*===================================*}
  1164 ML {*
  1165 *}
  1166 
  1167 section {*Integrate Program into Knowledge*}
  1168 ML {*
  1169 print_depth 999;
  1170 *}
  1171 
  1172 end
  1173