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