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