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