test/Tools/isac/ADDTESTS/course/SignalProcess/Build_Inverse_Z_Transform.thy
author Walther Neuper <neuper@ist.tugraz.at>
Mon, 05 Mar 2012 13:35:44 +0100
changeset 42384 cc5ba197f1b0
parent 42377 da3a55182442
child 42388 65da7356f5cd
permissions -rwxr-xr-x
made Test_Isac,thy run again

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