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