src/Tools/isac/Knowledge/Test.thy
author Walther Neuper <wneuper@ist.tugraz.at>
Fri, 18 May 2018 17:59:48 +0200
changeset 59430 b2345c1fe969
parent 59424 406681ebe781
child 59431 50918975b855
permissions -rw-r--r--
partial_function: clean Test.thy partially
wneuper@59430
     1
(* Knowledge for tests, specifically simplified or bound to a fixed state
wneuper@59430
     2
   for the purpose of simplifying tests.
neuper@37954
     3
   Author: Walther Neuper 2003
neuper@37954
     4
   (c) due to copyright terms
neuper@37954
     5
*) 
neuper@37906
     6
wneuper@59424
     7
theory Test imports Base_Tools Poly Rational Root Diff begin
neuper@37906
     8
 
wneuper@59430
     9
section {* consts *}
wneuper@59430
    10
subsection {* for scripts *}
neuper@37906
    11
consts
neuper@37906
    12
  Solve'_univar'_err
neuper@37954
    13
             :: "[bool,real,bool,  
neuper@37954
    14
		    bool list] => bool list"
wneuper@59334
    15
               ("((Script Solve'_univar'_err (_ _ _ =))// (_))" 9)
neuper@37906
    16
  Solve'_linear
neuper@37954
    17
             :: "[bool,real,  
neuper@37954
    18
		    bool list] => bool list"
wneuper@59334
    19
               ("((Script Solve'_linear (_ _ =))// (_))" 9)
wneuper@59430
    20
  Solve'_root'_equation 
wneuper@59430
    21
             :: "[bool,real,  
wneuper@59430
    22
		    bool list] => bool list"
wneuper@59430
    23
               ("((Script Solve'_root'_equation (_ _ =))// (_))" 9)
wneuper@59430
    24
  Solve'_plain'_square 
wneuper@59430
    25
             :: "[bool,real,  
wneuper@59430
    26
		    bool list] => bool list"
wneuper@59430
    27
               ("((Script Solve'_plain'_square (_ _ =))// (_))" 9)
wneuper@59430
    28
  Norm'_univar'_equation 
wneuper@59430
    29
             :: "[bool,real,  
wneuper@59430
    30
		    bool] => bool"
wneuper@59430
    31
               ("((Script Norm'_univar'_equation (_ _ =))// (_))" 9)
wneuper@59430
    32
  STest'_simplify
wneuper@59430
    33
             :: "['z,  
wneuper@59430
    34
		    'z] => 'z"
wneuper@59430
    35
               ("((Script STest'_simplify (_ =))// (_))" 9)
neuper@37906
    36
neuper@37906
    37
wneuper@59430
    38
subsection {* for problems *}
wneuper@59430
    39
consts
neuper@42008
    40
  "is'_root'_free"   :: "'a => bool"      ("is'_root'_free _" 10)
neuper@42008
    41
  "contains'_root"   :: "'a => bool"      ("contains'_root _" 10)
neuper@42008
    42
neuper@42008
    43
  "precond'_rootmet" :: "'a => bool"      ("precond'_rootmet _" 10)
neuper@42008
    44
  "precond'_rootpbl" :: "'a => bool"      ("precond'_rootpbl _" 10)
neuper@42008
    45
  "precond'_submet"  :: "'a => bool"      ("precond'_submet _" 10)
neuper@42008
    46
  "precond'_subpbl"  :: "'a => bool"      ("precond'_subpbl _" 10)
neuper@37906
    47
neuper@37906
    48
wneuper@59430
    49
section {* functions *}
wneuper@59430
    50
ML {*
wneuper@59430
    51
fun bin_o (Const (op_, (Type ("fun", [Type (s2, []), Type ("fun", [Type (s4, _),Type (s5, _)])]))))
wneuper@59430
    52
      = if s2 = s4 andalso s4 = s5 then [op_] else []
wneuper@59430
    53
    | bin_o _ = [];
neuper@37906
    54
wneuper@59430
    55
fun bin_op (t1 $ t2) = union op = (bin_op t1) (bin_op t2)
wneuper@59430
    56
  | bin_op t         =  bin_o t;
wneuper@59430
    57
fun is_bin_op t = (bin_op t <> []);
neuper@37906
    58
wneuper@59430
    59
fun bin_op_arg1 ((Const (_,
wneuper@59430
    60
    (Type ("fun", [Type (_, []), Type ("fun", [Type _, Type _])])))) $ arg1 $ _) = arg1
wneuper@59430
    61
  | bin_op_arg1 t = raise ERROR ("bin_op_arg1: t = " ^ Rule.term2str t);
wneuper@59430
    62
fun bin_op_arg2 ((Const (_,
wneuper@59430
    63
    (Type ("fun", [Type (_, []),Type ("fun", [Type _, Type _])]))))$ _ $ arg2) = arg2
wneuper@59430
    64
  | bin_op_arg2 t = raise ERROR ("bin_op_arg1: t = " ^ Rule.term2str t);
neuper@37906
    65
wneuper@59430
    66
exception NO_EQUATION_TERM;
wneuper@59430
    67
fun is_equation ((Const ("HOL.eq",
wneuper@59430
    68
    (Type ("fun", [Type (_, []), Type ("fun", [Type (_, []),Type ("bool",[])])])))) $ _ $ _) = true
wneuper@59430
    69
  | is_equation _ = false;
wneuper@59430
    70
fun equ_lhs ((Const ("HOL.eq",
wneuper@59430
    71
    (Type ("fun", [Type (_, []), Type ("fun", [Type (_, []),Type ("bool",[])])])))) $ l $ _) = l
wneuper@59430
    72
  | equ_lhs _ = raise NO_EQUATION_TERM;
wneuper@59430
    73
fun equ_rhs ((Const ("HOL.eq", (Type ("fun",
wneuper@59430
    74
		 [Type (_, []), Type ("fun", [Type (_, []), Type ("bool",[])])])))) $ _ $ r) = r
wneuper@59430
    75
  | equ_rhs _ = raise NO_EQUATION_TERM;
neuper@37906
    76
wneuper@59430
    77
fun atom (Const (_, Type (_,[]))) = true
wneuper@59430
    78
  | atom (Free (_, Type (_,[]))) = true
wneuper@59430
    79
  | atom (Var (_, Type (_,[]))) = true
wneuper@59430
    80
  | atom((Const ("Bin.integ_of_bin",_)) $ _) = true
wneuper@59430
    81
  | atom _ = false;
wneuper@59430
    82
wneuper@59430
    83
fun varids (Const (s, Type (_,[]))) = [strip_thy s]
wneuper@59430
    84
  | varids (Free (s, Type (_,[]))) = if TermC.is_num' s then [] else [strip_thy s]  
wneuper@59430
    85
  | varids (Var((s, _),Type (_,[]))) = [strip_thy s]
wneuper@59430
    86
(*| varids (_      (s,"?DUMMY"   )) =   ..ML-error *)
wneuper@59430
    87
  | varids((Const ("Bin.integ_of_bin",_)) $ _)= [](*8.01: superfluous?*)
wneuper@59430
    88
  | varids (Abs (a, _, t)) = union op = [a] (varids t)
wneuper@59430
    89
  | varids (t1 $ t2) = union op = (varids t1) (varids t2)
wneuper@59430
    90
  | varids _ = [];
wneuper@59430
    91
wneuper@59430
    92
fun bin_ops_only ((Const op_) $ t1 $ t2) =
wneuper@59430
    93
    if is_bin_op (Const op_) then bin_ops_only t1 andalso bin_ops_only t2 else false
wneuper@59430
    94
  | bin_ops_only t = if atom t then true else bin_ops_only t;
wneuper@59430
    95
wneuper@59430
    96
fun polynomial opl t _(* bdVar TODO *) =
wneuper@59430
    97
    subset op = (bin_op t, opl) andalso (bin_ops_only t);
wneuper@59430
    98
wneuper@59430
    99
fun poly_equ opl bdVar t = is_equation t (* bdVar TODO *) 
wneuper@59430
   100
    andalso polynomial opl (equ_lhs t) bdVar 
wneuper@59430
   101
    andalso polynomial opl (equ_rhs t) bdVar
wneuper@59430
   102
    andalso (subset op = (varids bdVar, varids (equ_lhs t)) orelse
wneuper@59430
   103
             subset op = (varids bdVar, varids (equ_lhs t)));
wneuper@59430
   104
wneuper@59430
   105
fun max (a,b) = if a < b then b else a;
wneuper@59430
   106
wneuper@59430
   107
fun degree addl mul bdVar t =
wneuper@59430
   108
let
wneuper@59430
   109
fun deg _ _ v (Const  (s, Type (_, []))) = if v=strip_thy s then 1 else 0
wneuper@59430
   110
  | deg _ _ v (Free   (s, Type (_, []))) = if v=strip_thy s then 1 else 0
wneuper@59430
   111
  | deg _ _ v (Var((s, _), Type (_, []))) = if v=strip_thy s then 1 else 0
wneuper@59430
   112
(*| deg _ _ v (_     (s,"?DUMMY"   ))          =   ..ML-error *) 
wneuper@59430
   113
  | deg _ _ _ ((Const ("Bin.integ_of_bin", _)) $ _ ) = 0 
wneuper@59430
   114
  | deg addl mul v (h $ t1 $ t2) =
wneuper@59430
   115
    if subset op = (bin_op h, addl)
wneuper@59430
   116
    then max (deg addl mul v t1  , deg addl mul v t2)
wneuper@59430
   117
    else (*mul!*)(deg addl mul v t1) + (deg addl mul v t2)
wneuper@59430
   118
  | deg _ _ _ t = raise ERROR ("deg: t = " ^ Rule.term2str t)
wneuper@59430
   119
in
wneuper@59430
   120
  if polynomial (addl @ [mul]) t bdVar
wneuper@59430
   121
  then SOME (deg addl mul (id_of bdVar) t)
wneuper@59430
   122
  else (NONE:int option)
wneuper@59430
   123
end;
wneuper@59430
   124
fun degree_ addl mul bdVar t = (* do not export *)
wneuper@59430
   125
let
wneuper@59430
   126
  fun opt (SOME i)= i
wneuper@59430
   127
	  | opt  NONE = 0
wneuper@59430
   128
in opt (degree addl mul bdVar t) end;
wneuper@59430
   129
wneuper@59430
   130
fun linear addl mul t bdVar = (degree_ addl mul bdVar t) < 2;
wneuper@59430
   131
wneuper@59430
   132
fun linear_equ addl mul bdVar t =
wneuper@59430
   133
  if is_equation t 
wneuper@59430
   134
  then
wneuper@59430
   135
    let
wneuper@59430
   136
      val degl = degree_ addl mul bdVar (equ_lhs t);
wneuper@59430
   137
	    val degr = degree_ addl mul bdVar (equ_rhs t)
wneuper@59430
   138
	  in if (degl>0 orelse degr>0)andalso max(degl,degr) < 2 then true else false end
wneuper@59430
   139
  else false;
wneuper@59430
   140
(* strip_thy op_  before *)
wneuper@59430
   141
fun is_div_op (dv, (Const (op_,
wneuper@59430
   142
    (Type ("fun", [Type (_, []), Type ("fun", [Type _, Type _])])))) ) = (dv = strip_thy op_)
wneuper@59430
   143
  | is_div_op _ = false;
wneuper@59430
   144
wneuper@59430
   145
fun is_denom bdVar div_op t =
wneuper@59430
   146
    let fun is bool[v]dv (Const  (s,Type(_,[])))= bool andalso(if v=strip_thy s then true else false)
wneuper@59430
   147
	  | is bool[v]dv (Free   (s,Type(_,[])))= bool andalso(if v=strip_thy s then true else false) 
wneuper@59430
   148
	  | is bool[v]dv (Var((s,_),Type(_,[])))= bool andalso(if v=strip_thy s then true else false)
wneuper@59430
   149
	  | is bool[v]dv((Const ("Bin.integ_of_bin",_)) $ _) = false
wneuper@59430
   150
	  | is bool[v]dv (h$n$d) = 
wneuper@59430
   151
	      if is_div_op(dv,h) 
wneuper@59430
   152
	      then (is false[v]dv n)orelse(is true[v]dv d)
wneuper@59430
   153
	      else (is bool [v]dv n)orelse(is bool[v]dv d)
wneuper@59430
   154
in is false (varids bdVar) (strip_thy div_op) t end;
wneuper@59430
   155
wneuper@59430
   156
wneuper@59430
   157
fun rational t div_op bdVar = 
wneuper@59430
   158
    is_denom bdVar div_op t andalso bin_ops_only t;
wneuper@59430
   159
wneuper@59430
   160
*}
wneuper@59430
   161
wneuper@59430
   162
section {* axiomatizations *}
neuper@52148
   163
axiomatization where (*TODO: prove as theorems*)
neuper@37906
   164
neuper@52148
   165
  radd_mult_distrib2:      "(k::real) * (m + n) = k * m + k * n" and
neuper@52148
   166
  rdistr_right_assoc:      "(k::real) + l * n + m * n = k + (l + m) * n" and
neuper@52148
   167
  rdistr_right_assoc_p:    "l * n + (m * n + (k::real)) = (l + m) * n + k" and
neuper@52148
   168
  rdistr_div_right:        "((k::real) + l) / n = k / n + l / n" and
neuper@37983
   169
  rcollect_right:
neuper@52148
   170
          "[| l is_const; m is_const |] ==> (l::real)*n + m*n = (l + m) * n" and
neuper@37983
   171
  rcollect_one_left:
neuper@52148
   172
          "m is_const ==> (n::real) + m * n = (1 + m) * n" and
neuper@37983
   173
  rcollect_one_left_assoc:
neuper@52148
   174
          "m is_const ==> (k::real) + n + m * n = k + (1 + m) * n" and
neuper@37983
   175
  rcollect_one_left_assoc_p:
neuper@52148
   176
          "m is_const ==> n + (m * n + (k::real)) = (1 + m) * n + k" and
neuper@37906
   177
neuper@52148
   178
  rtwo_of_the_same:        "a + a = 2 * a" and
neuper@52148
   179
  rtwo_of_the_same_assoc:  "(x + a) + a = x + 2 * a" and
neuper@52148
   180
  rtwo_of_the_same_assoc_p:"a + (a + x) = 2 * a + x" and
neuper@37906
   181
neuper@52148
   182
  rcancel_den:             "not(a=0) ==> a * (b / a) = b" and
neuper@52148
   183
  rcancel_const:           "[| a is_const; b is_const |] ==> a*(x/b) = a/b*x" and
neuper@52148
   184
  rshift_nominator:        "(a::real) * b / c = a / c * b" and
neuper@37906
   185
neuper@52148
   186
  exp_pow:                 "(a ^^^ b) ^^^ c = a ^^^ (b * c)" and
neuper@52148
   187
  rsqare:                  "(a::real) * a = a ^^^ 2" and
neuper@52148
   188
  power_1:                 "(a::real) ^^^ 1 = a" and
neuper@52148
   189
  rbinom_power_2:          "((a::real) + b)^^^ 2 = a^^^ 2 + 2*a*b + b^^^ 2" and
neuper@37906
   190
neuper@52148
   191
  rmult_1:                 "1 * k = (k::real)" and
neuper@52148
   192
  rmult_1_right:           "k * 1 = (k::real)" and
neuper@52148
   193
  rmult_0:                 "0 * k = (0::real)" and
neuper@52148
   194
  rmult_0_right:           "k * 0 = (0::real)" and
neuper@52148
   195
  radd_0:                  "0 + k = (k::real)" and
neuper@52148
   196
  radd_0_right:            "k + 0 = (k::real)" and
neuper@37906
   197
neuper@37983
   198
  radd_real_const_eq:
neuper@52148
   199
          "[| a is_const; c is_const; d is_const |] ==> a/d + c/d = (a+c)/(d::real)" and
neuper@37983
   200
  radd_real_const:
neuper@37906
   201
          "[| a is_const; b is_const; c is_const; d is_const |] ==> a/b + c/d = (a*d + b*c)/(b*(d::real))"  
neuper@52148
   202
   and
neuper@37906
   203
(*for AC-operators*)
neuper@52148
   204
  radd_commute:            "(m::real) + (n::real) = n + m" and
neuper@52148
   205
  radd_left_commute:       "(x::real) + (y + z) = y + (x + z)" and
neuper@52148
   206
  radd_assoc:              "(m::real) + n + k = m + (n + k)" and
neuper@52148
   207
  rmult_commute:           "(m::real) * n = n * m" and
neuper@52148
   208
  rmult_left_commute:      "(x::real) * (y * z) = y * (x * z)" and
neuper@52148
   209
  rmult_assoc:             "(m::real) * n * k = m * (n * k)" and
neuper@37906
   210
neuper@37906
   211
(*for equations: 'bdv' is a meta-constant*)
neuper@52148
   212
  risolate_bdv_add:       "((k::real) + bdv = m) = (bdv = m + (-1)*k)" and
neuper@52148
   213
  risolate_bdv_mult_add:  "((k::real) + n*bdv = m) = (n*bdv = m + (-1)*k)" and
neuper@52148
   214
  risolate_bdv_mult:      "((n::real) * bdv = m) = (bdv = m / n)" and
neuper@37906
   215
neuper@37983
   216
  rnorm_equation_add:
neuper@52148
   217
      "~(b =!= 0) ==> (a = b) = (a + (-1)*b = 0)" and
neuper@37906
   218
neuper@37906
   219
(*17.9.02 aus SqRoot.thy------------------------------vvv---*) 
neuper@52148
   220
  root_ge0:            "0 <= a ==> 0 <= sqrt a" and
neuper@37906
   221
  (*should be dropped with better simplification in eval_rls ...*)
neuper@37983
   222
  root_add_ge0:
neuper@52148
   223
	"[| 0 <= a; 0 <= b |] ==> (0 <= sqrt a + sqrt b) = True" and
neuper@37983
   224
  root_ge0_1:
neuper@52148
   225
	"[| 0<=a; 0<=b; 0<=c |] ==> (0 <= a * sqrt b + sqrt c) = True" and
neuper@37983
   226
  root_ge0_2:
neuper@52148
   227
	"[| 0<=a; 0<=b; 0<=c |] ==> (0 <= sqrt a + b * sqrt c) = True" and
neuper@37906
   228
neuper@37906
   229
neuper@52148
   230
  rroot_square_inv:         "(sqrt a)^^^ 2 = a" and
neuper@52148
   231
  rroot_times_root:         "sqrt a * sqrt b = sqrt(a*b)" and
neuper@52148
   232
  rroot_times_root_assoc:   "(a * sqrt b) * sqrt c = a * sqrt(b*c)" and
neuper@52148
   233
  rroot_times_root_assoc_p: "sqrt b * (sqrt c * a)= sqrt(b*c) * a" and
neuper@37906
   234
neuper@37906
   235
neuper@37906
   236
(*for root-equations*)
neuper@37983
   237
  square_equation_left:
neuper@52148
   238
          "[| 0 <= a; 0 <= b |] ==> (((sqrt a)=b)=(a=(b^^^ 2)))" and
neuper@37983
   239
  square_equation_right:
neuper@52148
   240
          "[| 0 <= a; 0 <= b |] ==> ((a=(sqrt b))=((a^^^ 2)=b))" and
neuper@37906
   241
  (*causes frequently non-termination:*)
neuper@37983
   242
  square_equation:  
neuper@52148
   243
          "[| 0 <= a; 0 <= b |] ==> ((a=b)=((a^^^ 2)=b^^^ 2))" and
neuper@37906
   244
  
neuper@52148
   245
  risolate_root_add:        "(a+  sqrt c = d) = (  sqrt c = d + (-1)*a)" and
neuper@52148
   246
  risolate_root_mult:       "(a+b*sqrt c = d) = (b*sqrt c = d + (-1)*a)" and
neuper@52148
   247
  risolate_root_div:        "(a * sqrt c = d) = (  sqrt c = d / a)" and
neuper@37906
   248
neuper@37906
   249
(*for polynomial equations of degree 2; linear case in RatArith*)
neuper@52148
   250
  mult_square:		"(a*bdv^^^2 = b) = (bdv^^^2 = b / a)" and
neuper@52148
   251
  constant_square:       "(a + bdv^^^2 = b) = (bdv^^^2 = b + -1*a)" and
neuper@52148
   252
  constant_mult_square:  "(a + b*bdv^^^2 = c) = (b*bdv^^^2 = c + -1*a)" and
neuper@37906
   253
neuper@37983
   254
  square_equality: 
neuper@52148
   255
	     "0 <= a ==> (x^^^2 = a) = ((x=sqrt a) | (x=-1*sqrt a))" and
neuper@37983
   256
  square_equality_0:
neuper@52148
   257
	     "(x^^^2 = 0) = (x = 0)" and
neuper@37906
   258
neuper@37906
   259
(*isolate root on the LEFT hand side of the equation
neuper@37906
   260
  otherwise shuffling from left to right would not terminate*)  
neuper@37906
   261
neuper@37983
   262
  rroot_to_lhs:
neuper@52148
   263
          "is_root_free a ==> (a = sqrt b) = (a + (-1)*sqrt b = 0)" and
neuper@37983
   264
  rroot_to_lhs_mult:
neuper@52148
   265
          "is_root_free a ==> (a = c*sqrt b) = (a + (-1)*c*sqrt b = 0)" and
neuper@37983
   266
  rroot_to_lhs_add_mult:
neuper@37906
   267
          "is_root_free a ==> (a = d+c*sqrt b) = (a + (-1)*c*sqrt b = d)"
neuper@37906
   268
(*17.9.02 aus SqRoot.thy------------------------------^^^---*)  
neuper@37906
   269
wneuper@59430
   270
section {* eval functions *}
neuper@37954
   271
ML {*
neuper@37972
   272
val thy = @{theory};
neuper@37972
   273
neuper@37954
   274
(** evaluation of numerals and predicates **)
neuper@37954
   275
neuper@37954
   276
(*does a term contain a root ?*)
neuper@37954
   277
fun eval_contains_root (thmid:string) _ 
neuper@37954
   278
		       (t as (Const("Test.contains'_root",t0) $ arg)) thy = 
neuper@38053
   279
  if member op = (ids_of arg) "sqrt"
wneuper@59416
   280
  then SOME (TermC.mk_thmid thmid (Rule.term_to_string''' thy arg)"",
wneuper@59390
   281
	       HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True})))
wneuper@59416
   282
  else SOME (TermC.mk_thmid thmid (Rule.term_to_string''' thy arg)"",
wneuper@59390
   283
	       HOLogic.Trueprop $ (TermC.mk_equality (t, @{term False})))
neuper@38053
   284
| eval_contains_root _ _ _ _ = NONE; 
neuper@42008
   285
neuper@42008
   286
(*dummy precondition for root-met of x+1=2*)
neuper@42008
   287
fun eval_precond_rootmet (thmid:string) _ (t as (Const ("Test.precond'_rootmet", _) $ arg)) thy = 
wneuper@59416
   288
    SOME (TermC.mk_thmid thmid (Rule.term_to_string''' thy arg)"",
wneuper@59390
   289
      HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True})))
neuper@52070
   290
  | eval_precond_rootmet _ _ _ _ = NONE; 
neuper@42008
   291
neuper@42008
   292
(*dummy precondition for root-pbl of x+1=2*)
neuper@42008
   293
fun eval_precond_rootpbl (thmid:string) _ (t as (Const ("Test.precond'_rootpbl", _) $ arg)) thy = 
wneuper@59416
   294
    SOME (TermC.mk_thmid thmid (Rule.term_to_string''' thy arg) "",
wneuper@59390
   295
	    HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True})))
s1210629013@52159
   296
	| eval_precond_rootpbl _ _ _ _ = NONE;
s1210629013@52145
   297
*}
s1210629013@52145
   298
setup {* KEStore_Elems.add_calcs
wneuper@59403
   299
  [("contains_root", ("Test.contains'_root", eval_contains_root"#contains_root_")),
s1210629013@52145
   300
    ("Test.precond'_rootmet", ("Test.precond'_rootmet", eval_precond_rootmet"#Test.precond_rootmet_")),
s1210629013@52145
   301
    ("Test.precond'_rootpbl", ("Test.precond'_rootpbl",
wneuper@59430
   302
        eval_precond_rootpbl"#Test.precond_rootpbl_"))]
wneuper@59430
   303
*}
s1210629013@52145
   304
ML {*
wneuper@59430
   305
(*8.4.03  aus Poly.ML--------------------------------vvv---
wneuper@59430
   306
  make_polynomial  ---> make_poly
wneuper@59430
   307
  ^-- for user          ^-- for systest _ONLY_*)  
wneuper@59430
   308
wneuper@59430
   309
local (*. for make_polytest .*)
wneuper@59430
   310
wneuper@59430
   311
open Term;  (* for type order = EQUAL | LESS | GREATER *)
wneuper@59430
   312
wneuper@59430
   313
fun pr_ord EQUAL = "EQUAL"
wneuper@59430
   314
  | pr_ord LESS  = "LESS"
wneuper@59430
   315
  | pr_ord GREATER = "GREATER";
wneuper@59430
   316
wneuper@59430
   317
fun dest_hd' (Const (a, T)) =                          (* ~ term.ML *)
wneuper@59430
   318
  (case a of
wneuper@59430
   319
     "Atools.pow" => ((("|||||||||||||", 0), T), 0)           (*WN greatest *)
wneuper@59430
   320
   | _ => (((a, 0), T), 0))
wneuper@59430
   321
  | dest_hd' (Free (a, T)) = (((a, 0), T), 1)
wneuper@59430
   322
  | dest_hd' (Var v) = (v, 2)
wneuper@59430
   323
  | dest_hd' (Bound i) = ((("", i), dummyT), 3)
wneuper@59430
   324
  | dest_hd' (Abs (_, T, _)) = ((("", 0), T), 4);
wneuper@59430
   325
(* RL *)
wneuper@59430
   326
fun get_order_pow (t $ (Free(order,_))) = 
wneuper@59430
   327
    	(case TermC.int_of_str_opt (order) of
wneuper@59430
   328
	             SOME d => d
wneuper@59430
   329
		   | NONE   => 0)
wneuper@59430
   330
  | get_order_pow _ = 0;
wneuper@59430
   331
wneuper@59430
   332
fun size_of_term' (Const(str,_) $ t) =
wneuper@59430
   333
  if "Atools.pow"=str then 1000 + size_of_term' t else 1 + size_of_term' t(*WN*)
wneuper@59430
   334
  | size_of_term' (Abs (_,_,body)) = 1 + size_of_term' body
wneuper@59430
   335
  | size_of_term' (f$t) = size_of_term' f  +  size_of_term' t
wneuper@59430
   336
  | size_of_term' _ = 1;
wneuper@59430
   337
fun term_ord' pr thy (Abs (_, T, t), Abs(_, U, u)) =       (* ~ term.ML *)
wneuper@59430
   338
    (case term_ord' pr thy (t, u) of EQUAL => Term_Ord.typ_ord (T, U) 
wneuper@59430
   339
                                   | ord => ord)
wneuper@59430
   340
  | term_ord' pr thy (t, u) =
wneuper@59430
   341
    (if pr then 
wneuper@59430
   342
	 let val (f, ts) = strip_comb t and (g, us) = strip_comb u;
wneuper@59430
   343
	     val _ = tracing ("t= f@ts= \"" ^ Rule.term2str f ^ "\" @ \"[" ^
wneuper@59430
   344
	                      commas(map Rule.term2str ts) ^ "]\"")
wneuper@59430
   345
	     val _ = tracing ("u= g@us= \"" ^ Rule.term2str g ^"\" @ \"[" ^
wneuper@59430
   346
	                      commas(map Rule.term2str us) ^"]\"")
wneuper@59430
   347
	     val _ = tracing ("size_of_term(t,u)= (" ^
wneuper@59430
   348
	                      string_of_int (size_of_term' t) ^ ", " ^
wneuper@59430
   349
	                      string_of_int (size_of_term' u) ^ ")")
wneuper@59430
   350
	     val _ = tracing ("hd_ord(f,g)      = " ^ (pr_ord o hd_ord) (f,g))
wneuper@59430
   351
	     val _ = tracing ("terms_ord(ts,us) = " ^
wneuper@59430
   352
			      (pr_ord o terms_ord str false) (ts,us));
wneuper@59430
   353
	     val _ = tracing "-------"
wneuper@59430
   354
	 in () end
wneuper@59430
   355
       else ();
wneuper@59430
   356
	 case int_ord (size_of_term' t, size_of_term' u) of
wneuper@59430
   357
	   EQUAL =>
wneuper@59430
   358
	     let val (f, ts) = strip_comb t and (g, us) = strip_comb u in
wneuper@59430
   359
	       (case hd_ord (f, g) of EQUAL => (terms_ord str pr) (ts, us) 
wneuper@59430
   360
	     | ord => ord)
wneuper@59430
   361
	     end
wneuper@59430
   362
	 | ord => ord)
wneuper@59430
   363
and hd_ord (f, g) =                                        (* ~ term.ML *)
wneuper@59430
   364
  prod_ord (prod_ord Term_Ord.indexname_ord Term_Ord.typ_ord) int_ord (dest_hd' f, dest_hd' g)
wneuper@59430
   365
and terms_ord str pr (ts, us) = 
wneuper@59430
   366
    list_ord (term_ord' pr (Celem.assoc_thy "Isac"))(ts, us);
wneuper@59430
   367
in
wneuper@59430
   368
wneuper@59430
   369
fun ord_make_polytest (pr:bool) thy (_: Rule.subst) tu = 
wneuper@59430
   370
    (term_ord' pr thy(***) tu = LESS );
wneuper@59430
   371
wneuper@59430
   372
end;(*local*)
wneuper@59430
   373
*} 
wneuper@59430
   374
wneuper@59430
   375
section {* term order *}
wneuper@59430
   376
ML {*
wneuper@59416
   377
fun term_order (_: Rule.subst) tu = (term_ordI [] tu = LESS);
wneuper@59430
   378
*}
s1210629013@52145
   379
wneuper@59430
   380
section {* rulesets *}
wneuper@59430
   381
ML {*
neuper@37954
   382
val testerls = 
wneuper@59416
   383
  Rule.Rls {id = "testerls", preconds = [], rew_ord = ("termlessI",termlessI), 
wneuper@59416
   384
      erls = Rule.e_rls, srls = Rule.Erls, 
neuper@42451
   385
      calc = [], errpatts = [], 
wneuper@59416
   386
      rules = [Rule.Thm ("refl",TermC.num_str @{thm refl}),
wneuper@59416
   387
	       Rule.Thm ("order_refl",TermC.num_str @{thm order_refl}),
wneuper@59416
   388
	       Rule.Thm ("radd_left_cancel_le",TermC.num_str @{thm radd_left_cancel_le}),
wneuper@59416
   389
	       Rule.Thm ("not_true",TermC.num_str @{thm not_true}),
wneuper@59416
   390
	       Rule.Thm ("not_false",TermC.num_str @{thm not_false}),
wneuper@59416
   391
	       Rule.Thm ("and_true",TermC.num_str @{thm and_true}),
wneuper@59416
   392
	       Rule.Thm ("and_false",TermC.num_str @{thm and_false}),
wneuper@59416
   393
	       Rule.Thm ("or_true",TermC.num_str @{thm or_true}),
wneuper@59416
   394
	       Rule.Thm ("or_false",TermC.num_str @{thm or_false}),
wneuper@59416
   395
	       Rule.Thm ("and_commute",TermC.num_str @{thm and_commute}),
wneuper@59416
   396
	       Rule.Thm ("or_commute",TermC.num_str @{thm or_commute}),
neuper@37954
   397
wneuper@59416
   398
	       Rule.Calc ("Atools.is'_const",eval_const "#is_const_"),
wneuper@59416
   399
	       Rule.Calc ("Tools.matches",eval_matches ""),
neuper@37954
   400
    
wneuper@59416
   401
	       Rule.Calc ("Groups.plus_class.plus",eval_binop "#add_"),
wneuper@59416
   402
	       Rule.Calc ("Groups.times_class.times",eval_binop "#mult_"),
wneuper@59416
   403
	       Rule.Calc ("Atools.pow" ,eval_binop "#power_"),
neuper@37954
   404
		    
wneuper@59416
   405
	       Rule.Calc ("Orderings.ord_class.less",eval_equ "#less_"),
wneuper@59416
   406
	       Rule.Calc ("Orderings.ord_class.less_eq",eval_equ "#less_equal_"),
neuper@37954
   407
	     	    
wneuper@59416
   408
	       Rule.Calc ("Atools.ident",eval_ident "#ident_")],
wneuper@59416
   409
      scr = Rule.Prog ((Thm.term_of o the o (TermC.parse thy)) "empty_script")
wneuper@59406
   410
      };      
neuper@38001
   411
*}
neuper@38001
   412
ML {*
neuper@37954
   413
(*.for evaluation of conditions in rewrite rules.*)
neuper@37954
   414
(*FIXXXXXXME 10.8.02: handle like _simplify*)
neuper@37954
   415
val tval_rls =  
wneuper@59416
   416
  Rule.Rls{id = "tval_rls", preconds = [], 
wneuper@59397
   417
      rew_ord = ("sqrt_right",sqrt_right false @{theory "Pure"}), 
wneuper@59416
   418
      erls=testerls,srls = Rule.e_rls, 
neuper@42451
   419
      calc=[], errpatts = [],
wneuper@59416
   420
      rules = [Rule.Thm ("refl",TermC.num_str @{thm refl}),
wneuper@59416
   421
	       Rule.Thm ("order_refl",TermC.num_str @{thm order_refl}),
wneuper@59416
   422
	       Rule.Thm ("radd_left_cancel_le",TermC.num_str @{thm radd_left_cancel_le}),
wneuper@59416
   423
	       Rule.Thm ("not_true",TermC.num_str @{thm not_true}),
wneuper@59416
   424
	       Rule.Thm ("not_false",TermC.num_str @{thm not_false}),
wneuper@59416
   425
	       Rule.Thm ("and_true",TermC.num_str @{thm and_true}),
wneuper@59416
   426
	       Rule.Thm ("and_false",TermC.num_str @{thm and_false}),
wneuper@59416
   427
	       Rule.Thm ("or_true",TermC.num_str @{thm or_true}),
wneuper@59416
   428
	       Rule.Thm ("or_false",TermC.num_str @{thm or_false}),
wneuper@59416
   429
	       Rule.Thm ("and_commute",TermC.num_str @{thm and_commute}),
wneuper@59416
   430
	       Rule.Thm ("or_commute",TermC.num_str @{thm or_commute}),
neuper@37954
   431
wneuper@59416
   432
	       Rule.Thm ("real_diff_minus",TermC.num_str @{thm real_diff_minus}),
neuper@37954
   433
wneuper@59416
   434
	       Rule.Thm ("root_ge0",TermC.num_str @{thm root_ge0}),
wneuper@59416
   435
	       Rule.Thm ("root_add_ge0",TermC.num_str @{thm root_add_ge0}),
wneuper@59416
   436
	       Rule.Thm ("root_ge0_1",TermC.num_str @{thm root_ge0_1}),
wneuper@59416
   437
	       Rule.Thm ("root_ge0_2",TermC.num_str @{thm root_ge0_2}),
neuper@37954
   438
wneuper@59416
   439
	       Rule.Calc ("Atools.is'_const",eval_const "#is_const_"),
wneuper@59416
   440
	       Rule.Calc ("Test.contains'_root",eval_contains_root "#eval_contains_root"),
wneuper@59416
   441
	       Rule.Calc ("Tools.matches",eval_matches ""),
wneuper@59416
   442
	       Rule.Calc ("Test.contains'_root",
neuper@37954
   443
		     eval_contains_root"#contains_root_"),
neuper@37954
   444
    
wneuper@59416
   445
	       Rule.Calc ("Groups.plus_class.plus",eval_binop "#add_"),
wneuper@59416
   446
	       Rule.Calc ("Groups.times_class.times",eval_binop "#mult_"),
wneuper@59416
   447
	       Rule.Calc ("NthRoot.sqrt",eval_sqrt "#sqrt_"),
wneuper@59416
   448
	       Rule.Calc ("Atools.pow" ,eval_binop "#power_"),
neuper@37954
   449
		    
wneuper@59416
   450
	       Rule.Calc ("Orderings.ord_class.less",eval_equ "#less_"),
wneuper@59416
   451
	       Rule.Calc ("Orderings.ord_class.less_eq",eval_equ "#less_equal_"),
neuper@37954
   452
	     	    
wneuper@59416
   453
	       Rule.Calc ("Atools.ident",eval_ident "#ident_")],
wneuper@59416
   454
      scr = Rule.Prog ((Thm.term_of o the o (TermC.parse thy)) "empty_script")
wneuper@59406
   455
      };      
neuper@38001
   456
*}
s1210629013@55444
   457
setup {* KEStore_Elems.add_rlss [("testerls", (Context.theory_name @{theory}, prep_rls' testerls))] *}
neuper@52155
   458
neuper@38001
   459
ML {*
neuper@37954
   460
(*make () dissappear*)   
neuper@37954
   461
val rearrange_assoc =
wneuper@59416
   462
  Rule.Rls{id = "rearrange_assoc", preconds = [], 
wneuper@59416
   463
      rew_ord = ("e_rew_ord",Rule.e_rew_ord), 
wneuper@59416
   464
      erls = Rule.e_rls, srls = Rule.e_rls, calc = [], errpatts = [],
neuper@37954
   465
      rules = 
wneuper@59416
   466
      [Rule.Thm ("sym_add_assoc",TermC.num_str (@{thm add.assoc} RS @{thm sym})),
wneuper@59416
   467
       Rule.Thm ("sym_rmult_assoc",TermC.num_str (@{thm rmult_assoc} RS @{thm sym}))],
wneuper@59416
   468
      scr = Rule.Prog ((Thm.term_of o the o (TermC.parse thy)) "empty_script")
wneuper@59406
   469
      };      
neuper@37954
   470
neuper@37954
   471
val ac_plus_times =
wneuper@59416
   472
  Rule.Rls{id = "ac_plus_times", preconds = [], rew_ord = ("term_order",term_order),
wneuper@59416
   473
      erls = Rule.e_rls, srls = Rule.e_rls, calc = [], errpatts = [],
neuper@37954
   474
      rules = 
wneuper@59416
   475
      [Rule.Thm ("radd_commute",TermC.num_str @{thm radd_commute}),
wneuper@59416
   476
       Rule.Thm ("radd_left_commute",TermC.num_str @{thm radd_left_commute}),
wneuper@59416
   477
       Rule.Thm ("add_assoc",TermC.num_str @{thm add.assoc}),
wneuper@59416
   478
       Rule.Thm ("rmult_commute",TermC.num_str @{thm rmult_commute}),
wneuper@59416
   479
       Rule.Thm ("rmult_left_commute",TermC.num_str @{thm rmult_left_commute}),
wneuper@59416
   480
       Rule.Thm ("rmult_assoc",TermC.num_str @{thm rmult_assoc})],
wneuper@59416
   481
      scr = Rule.Prog ((Thm.term_of o the o (TermC.parse thy)) "empty_script")
wneuper@59406
   482
      };      
neuper@37954
   483
wneuper@59389
   484
(*todo: replace by Rewrite("rnorm_equation_add",TermC.num_str @{thm rnorm_equation_add)*)
neuper@37954
   485
val norm_equation =
wneuper@59416
   486
  Rule.Rls{id = "norm_equation", preconds = [], rew_ord = ("e_rew_ord",Rule.e_rew_ord),
wneuper@59416
   487
      erls = tval_rls, srls = Rule.e_rls, calc = [], errpatts = [],
wneuper@59416
   488
      rules = [Rule.Thm ("rnorm_equation_add",TermC.num_str @{thm rnorm_equation_add})
neuper@37954
   489
	       ],
wneuper@59416
   490
      scr = Rule.Prog ((Thm.term_of o the o (TermC.parse thy)) "empty_script")
wneuper@59406
   491
      };      
neuper@38001
   492
*}
neuper@38001
   493
ML {*
neuper@37954
   494
(* expects * distributed over + *)
neuper@37954
   495
val Test_simplify =
wneuper@59416
   496
  Rule.Rls{id = "Test_simplify", preconds = [], 
wneuper@59397
   497
      rew_ord = ("sqrt_right",sqrt_right false @{theory "Pure"}),
wneuper@59416
   498
      erls = tval_rls, srls = Rule.e_rls, 
s1210629013@55444
   499
      calc=[(*since 040209 filled by prep_rls'*)], errpatts = [],
neuper@37954
   500
      rules = [
wneuper@59416
   501
	       Rule.Thm ("real_diff_minus",TermC.num_str @{thm real_diff_minus}),
wneuper@59416
   502
	       Rule.Thm ("radd_mult_distrib2",TermC.num_str @{thm radd_mult_distrib2}),
wneuper@59416
   503
	       Rule.Thm ("rdistr_right_assoc",TermC.num_str @{thm rdistr_right_assoc}),
wneuper@59416
   504
	       Rule.Thm ("rdistr_right_assoc_p",TermC.num_str @{thm rdistr_right_assoc_p}),
wneuper@59416
   505
	       Rule.Thm ("rdistr_div_right",TermC.num_str @{thm rdistr_div_right}),
wneuper@59416
   506
	       Rule.Thm ("rbinom_power_2",TermC.num_str @{thm rbinom_power_2}),	       
neuper@37954
   507
wneuper@59416
   508
               Rule.Thm ("radd_commute",TermC.num_str @{thm radd_commute}), 
wneuper@59416
   509
	       Rule.Thm ("radd_left_commute",TermC.num_str @{thm radd_left_commute}),
wneuper@59416
   510
	       Rule.Thm ("add_assoc",TermC.num_str @{thm add.assoc}),
wneuper@59416
   511
	       Rule.Thm ("rmult_commute",TermC.num_str @{thm rmult_commute}),
wneuper@59416
   512
	       Rule.Thm ("rmult_left_commute",TermC.num_str @{thm rmult_left_commute}),
wneuper@59416
   513
	       Rule.Thm ("rmult_assoc",TermC.num_str @{thm rmult_assoc}),
neuper@37954
   514
wneuper@59416
   515
	       Rule.Thm ("radd_real_const_eq",TermC.num_str @{thm radd_real_const_eq}),
wneuper@59416
   516
	       Rule.Thm ("radd_real_const",TermC.num_str @{thm radd_real_const}),
neuper@37954
   517
	       (* these 2 rules are invers to distr_div_right wrt. termination.
neuper@37954
   518
		  thus they MUST be done IMMEDIATELY before calc *)
wneuper@59416
   519
	       Rule.Calc ("Groups.plus_class.plus", eval_binop "#add_"), 
wneuper@59416
   520
	       Rule.Calc ("Groups.times_class.times", eval_binop "#mult_"),
wneuper@59416
   521
	       Rule.Calc ("Rings.divide_class.divide", eval_cancel "#divide_e"),
wneuper@59416
   522
	       Rule.Calc ("Atools.pow", eval_binop "#power_"),
neuper@37954
   523
wneuper@59416
   524
	       Rule.Thm ("rcollect_right",TermC.num_str @{thm rcollect_right}),
wneuper@59416
   525
	       Rule.Thm ("rcollect_one_left",TermC.num_str @{thm rcollect_one_left}),
wneuper@59416
   526
	       Rule.Thm ("rcollect_one_left_assoc",TermC.num_str @{thm rcollect_one_left_assoc}),
wneuper@59416
   527
	       Rule.Thm ("rcollect_one_left_assoc_p",TermC.num_str @{thm rcollect_one_left_assoc_p}),
neuper@37954
   528
wneuper@59416
   529
	       Rule.Thm ("rshift_nominator",TermC.num_str @{thm rshift_nominator}),
wneuper@59416
   530
	       Rule.Thm ("rcancel_den",TermC.num_str @{thm rcancel_den}),
wneuper@59416
   531
	       Rule.Thm ("rroot_square_inv",TermC.num_str @{thm rroot_square_inv}),
wneuper@59416
   532
	       Rule.Thm ("rroot_times_root",TermC.num_str @{thm rroot_times_root}),
wneuper@59416
   533
	       Rule.Thm ("rroot_times_root_assoc_p",TermC.num_str @{thm rroot_times_root_assoc_p}),
wneuper@59416
   534
	       Rule.Thm ("rsqare",TermC.num_str @{thm rsqare}),
wneuper@59416
   535
	       Rule.Thm ("power_1",TermC.num_str @{thm power_1}),
wneuper@59416
   536
	       Rule.Thm ("rtwo_of_the_same",TermC.num_str @{thm rtwo_of_the_same}),
wneuper@59416
   537
	       Rule.Thm ("rtwo_of_the_same_assoc_p",TermC.num_str @{thm rtwo_of_the_same_assoc_p}),
neuper@37954
   538
wneuper@59416
   539
	       Rule.Thm ("rmult_1",TermC.num_str @{thm rmult_1}),
wneuper@59416
   540
	       Rule.Thm ("rmult_1_right",TermC.num_str @{thm rmult_1_right}),
wneuper@59416
   541
	       Rule.Thm ("rmult_0",TermC.num_str @{thm rmult_0}),
wneuper@59416
   542
	       Rule.Thm ("rmult_0_right",TermC.num_str @{thm rmult_0_right}),
wneuper@59416
   543
	       Rule.Thm ("radd_0",TermC.num_str @{thm radd_0}),
wneuper@59416
   544
	       Rule.Thm ("radd_0_right",TermC.num_str @{thm radd_0_right})
neuper@37954
   545
	       ],
wneuper@59416
   546
      scr = Rule.Prog ((Thm.term_of o the o (TermC.parse thy)) "empty_script")
s1210629013@55444
   547
		    (*since 040209 filled by prep_rls': STest_simplify*)
wneuper@59406
   548
      };      
neuper@38001
   549
*}
neuper@38001
   550
ML {*
neuper@37954
   551
(*isolate the root in a root-equation*)
neuper@37954
   552
val isolate_root =
wneuper@59416
   553
  Rule.Rls{id = "isolate_root", preconds = [], rew_ord = ("e_rew_ord",Rule.e_rew_ord), 
wneuper@59416
   554
      erls=tval_rls,srls = Rule.e_rls, calc=[], errpatts = [],
wneuper@59416
   555
      rules = [Rule.Thm ("rroot_to_lhs",TermC.num_str @{thm rroot_to_lhs}),
wneuper@59416
   556
	       Rule.Thm ("rroot_to_lhs_mult",TermC.num_str @{thm rroot_to_lhs_mult}),
wneuper@59416
   557
	       Rule.Thm ("rroot_to_lhs_add_mult",TermC.num_str @{thm rroot_to_lhs_add_mult}),
wneuper@59416
   558
	       Rule.Thm ("risolate_root_add",TermC.num_str @{thm risolate_root_add}),
wneuper@59416
   559
	       Rule.Thm ("risolate_root_mult",TermC.num_str @{thm risolate_root_mult}),
wneuper@59416
   560
	       Rule.Thm ("risolate_root_div",TermC.num_str @{thm risolate_root_div})       ],
wneuper@59416
   561
      scr = Rule.Prog ((Thm.term_of o the o (TermC.parse thy)) 
neuper@37954
   562
      "empty_script")
wneuper@59406
   563
      };
neuper@37954
   564
neuper@37954
   565
(*isolate the bound variable in an equation; 'bdv' is a meta-constant*)
neuper@37954
   566
val isolate_bdv =
wneuper@59416
   567
    Rule.Rls{id = "isolate_bdv", preconds = [], rew_ord = ("e_rew_ord",Rule.e_rew_ord),
wneuper@59416
   568
	erls=tval_rls,srls = Rule.e_rls, calc= [], errpatts = [],
neuper@37954
   569
	rules = 
wneuper@59416
   570
	[Rule.Thm ("risolate_bdv_add",TermC.num_str @{thm risolate_bdv_add}),
wneuper@59416
   571
	 Rule.Thm ("risolate_bdv_mult_add",TermC.num_str @{thm risolate_bdv_mult_add}),
wneuper@59416
   572
	 Rule.Thm ("risolate_bdv_mult",TermC.num_str @{thm risolate_bdv_mult}),
wneuper@59416
   573
	 Rule.Thm ("mult_square",TermC.num_str @{thm mult_square}),
wneuper@59416
   574
	 Rule.Thm ("constant_square",TermC.num_str @{thm constant_square}),
wneuper@59416
   575
	 Rule.Thm ("constant_mult_square",TermC.num_str @{thm constant_mult_square})
neuper@37954
   576
	 ],
wneuper@59416
   577
	scr = Rule.Prog ((Thm.term_of o the o (TermC.parse thy)) 
neuper@37954
   578
			  "empty_script")
wneuper@59406
   579
	};      
neuper@38001
   580
*}
wneuper@59374
   581
ML {* val prep_rls' = LTool.prep_rls @{theory}; *}
neuper@52125
   582
setup {* KEStore_Elems.add_rlss 
s1210629013@55444
   583
  [("Test_simplify", (Context.theory_name @{theory}, prep_rls' Test_simplify)), 
s1210629013@55444
   584
  ("tval_rls", (Context.theory_name @{theory}, prep_rls' tval_rls)), 
s1210629013@55444
   585
  ("isolate_root", (Context.theory_name @{theory}, prep_rls' isolate_root)), 
s1210629013@55444
   586
  ("isolate_bdv", (Context.theory_name @{theory}, prep_rls' isolate_bdv)), 
s1210629013@55444
   587
  ("matches", (Context.theory_name @{theory}, prep_rls'
wneuper@59430
   588
    (Rule.append_rls "matches" testerls [Rule.Calc ("Tools.matches",eval_matches "#matches_")])))]
wneuper@59430
   589
*}
neuper@37954
   590
wneuper@59430
   591
subsection {* problems *}
neuper@37954
   592
(** problem types **)
s1210629013@55359
   593
setup {* KEStore_Elems.add_pbts
wneuper@59416
   594
  [(Specify.prep_pbt thy "pbl_test" [] Celem.e_pblID (["test"], [], Rule.e_rls, NONE, [])),
wneuper@59406
   595
    (Specify.prep_pbt thy "pbl_test_equ" [] Celem.e_pblID
s1210629013@55339
   596
      (["equation","test"],
s1210629013@55339
   597
        [("#Given" ,["equality e_e","solveFor v_v"]),
s1210629013@55339
   598
           ("#Where" ,["matches (?a = ?b) e_e"]),
s1210629013@55339
   599
           ("#Find"  ,["solutions v_v'i'"])],
s1210629013@55339
   600
        assoc_rls' @{theory} "matches", SOME "solve (e_e::bool, v_v)", [])),
wneuper@59406
   601
    (Specify.prep_pbt thy "pbl_test_uni" [] Celem.e_pblID
s1210629013@55339
   602
      (["univariate","equation","test"],
s1210629013@55339
   603
        [("#Given" ,["equality e_e","solveFor v_v"]),
s1210629013@55339
   604
           ("#Where" ,["matches (?a = ?b) e_e"]),
s1210629013@55339
   605
           ("#Find"  ,["solutions v_v'i'"])],
s1210629013@55339
   606
        assoc_rls' @{theory} "matches", SOME "solve (e_e::bool, v_v)", [])),
wneuper@59406
   607
    (Specify.prep_pbt thy "pbl_test_uni_lin" [] Celem.e_pblID
s1210629013@55339
   608
      (["LINEAR","univariate","equation","test"],
s1210629013@55339
   609
        [("#Given" ,["equality e_e","solveFor v_v"]),
s1210629013@55339
   610
           ("#Where" ,["(matches (   v_v = 0) e_e) | (matches (   ?b*v_v = 0) e_e) |" ^
s1210629013@55339
   611
             "(matches (?a+v_v = 0) e_e) | (matches (?a+?b*v_v = 0) e_e)  "]),
s1210629013@55339
   612
           ("#Find"  ,["solutions v_v'i'"])],
s1210629013@55339
   613
        assoc_rls' @{theory} "matches", 
wneuper@59430
   614
        SOME "solve (e_e::bool, v_v)", [["Test","solve_linear"]]))]
neuper@38001
   615
*}
neuper@38001
   616
ML {*
wneuper@59416
   617
Rule.rew_ord' := overwritel (! Rule.rew_ord',
neuper@37954
   618
[("termlessI", termlessI),
neuper@37954
   619
 ("ord_make_polytest", ord_make_polytest false thy)
neuper@37954
   620
 ]);
neuper@37954
   621
neuper@37954
   622
val make_polytest =
wneuper@59416
   623
  Rule.Rls{id = "make_polytest", preconds = []:term list, 
neuper@52105
   624
      rew_ord = ("ord_make_polytest", ord_make_polytest false @{theory "Poly"}),
wneuper@59416
   625
      erls = testerls, srls = Rule.Erls,
neuper@38014
   626
      calc = [("PLUS"  , ("Groups.plus_class.plus", eval_binop "#add_")), 
neuper@38034
   627
	      ("TIMES" , ("Groups.times_class.times", eval_binop "#mult_")),
neuper@37954
   628
	      ("POWER", ("Atools.pow", eval_binop "#power_"))
neuper@42451
   629
	      ], errpatts = [],
wneuper@59416
   630
      rules = [Rule.Thm ("real_diff_minus",TermC.num_str @{thm real_diff_minus}),
neuper@37954
   631
	       (*"a - b = a + (-1) * b"*)
wneuper@59416
   632
	       Rule.Thm ("distrib_right" ,TermC.num_str @{thm distrib_right}),
neuper@37954
   633
	       (*"(z1.0 + z2.0) * w = z1.0 * w + z2.0 * w"*)
wneuper@59416
   634
	       Rule.Thm ("distrib_left",TermC.num_str @{thm distrib_left}),
neuper@37954
   635
	       (*"w * (z1.0 + z2.0) = w * z1.0 + w * z2.0"*)
wneuper@59416
   636
	       Rule.Thm ("left_diff_distrib" ,TermC.num_str @{thm left_diff_distrib}),
neuper@37954
   637
	       (*"(z1.0 - z2.0) * w = z1.0 * w - z2.0 * w"*)
wneuper@59416
   638
	       Rule.Thm ("right_diff_distrib",TermC.num_str @{thm right_diff_distrib}),
neuper@37954
   639
	       (*"w * (z1.0 - z2.0) = w * z1.0 - w * z2.0"*)
wneuper@59416
   640
	       Rule.Thm ("mult_1_left",TermC.num_str @{thm mult_1_left}),                 
neuper@37954
   641
	       (*"1 * z = z"*)
wneuper@59416
   642
	       Rule.Thm ("mult_zero_left",TermC.num_str @{thm mult_zero_left}),        
neuper@37954
   643
	       (*"0 * z = 0"*)
wneuper@59416
   644
	       Rule.Thm ("add_0_left",TermC.num_str @{thm add_0_left}),
neuper@37954
   645
	       (*"0 + z = z"*)
neuper@37954
   646
neuper@37954
   647
	       (*AC-rewriting*)
wneuper@59416
   648
	       Rule.Thm ("mult_commute",TermC.num_str @{thm mult.commute}),
neuper@37954
   649
	       (* z * w = w * z *)
wneuper@59416
   650
	       Rule.Thm ("real_mult_left_commute",TermC.num_str @{thm real_mult_left_commute}),
neuper@37954
   651
	       (*z1.0 * (z2.0 * z3.0) = z2.0 * (z1.0 * z3.0)*)
wneuper@59416
   652
	       Rule.Thm ("mult_assoc",TermC.num_str @{thm mult.assoc}),		
neuper@37954
   653
	       (*z1.0 * z2.0 * z3.0 = z1.0 * (z2.0 * z3.0)*)
wneuper@59416
   654
	       Rule.Thm ("add_commute",TermC.num_str @{thm add.commute}),	
neuper@37954
   655
	       (*z + w = w + z*)
wneuper@59416
   656
	       Rule.Thm ("add_left_commute",TermC.num_str @{thm add.left_commute}),
neuper@37954
   657
	       (*x + (y + z) = y + (x + z)*)
wneuper@59416
   658
	       Rule.Thm ("add_assoc",TermC.num_str @{thm add.assoc}),	               
neuper@37954
   659
	       (*z1.0 + z2.0 + z3.0 = z1.0 + (z2.0 + z3.0)*)
neuper@37954
   660
wneuper@59416
   661
	       Rule.Thm ("sym_realpow_twoI",
wneuper@59389
   662
                     TermC.num_str (@{thm realpow_twoI} RS @{thm sym})),	
neuper@37954
   663
	       (*"r1 * r1 = r1 ^^^ 2"*)
wneuper@59416
   664
	       Rule.Thm ("realpow_plus_1",TermC.num_str @{thm realpow_plus_1}),		
neuper@37954
   665
	       (*"r * r ^^^ n = r ^^^ (n + 1)"*)
wneuper@59416
   666
	       Rule.Thm ("sym_real_mult_2",
wneuper@59389
   667
                     TermC.num_str (@{thm real_mult_2} RS @{thm sym})),	
neuper@37954
   668
	       (*"z1 + z1 = 2 * z1"*)
wneuper@59416
   669
	       Rule.Thm ("real_mult_2_assoc",TermC.num_str @{thm real_mult_2_assoc}),	
neuper@37954
   670
	       (*"z1 + (z1 + k) = 2 * z1 + k"*)
neuper@37954
   671
wneuper@59416
   672
	       Rule.Thm ("real_num_collect",TermC.num_str @{thm real_num_collect}), 
neuper@37954
   673
	       (*"[| l is_const; m is_const |]==>l * n + m * n = (l + m) * n"*)
wneuper@59416
   674
	       Rule.Thm ("real_num_collect_assoc",TermC.num_str @{thm real_num_collect_assoc}),
neuper@37954
   675
	       (*"[| l is_const; m is_const |] ==>  
neuper@37954
   676
				l * n + (m * n + k) =  (l + m) * n + k"*)
wneuper@59416
   677
	       Rule.Thm ("real_one_collect",TermC.num_str @{thm real_one_collect}),	
neuper@37954
   678
	       (*"m is_const ==> n + m * n = (1 + m) * n"*)
wneuper@59416
   679
	       Rule.Thm ("real_one_collect_assoc",TermC.num_str @{thm real_one_collect_assoc}), 
neuper@37954
   680
	       (*"m is_const ==> k + (n + m * n) = k + (1 + m) * n"*)
neuper@37954
   681
wneuper@59416
   682
	       Rule.Calc ("Groups.plus_class.plus", eval_binop "#add_"), 
wneuper@59416
   683
	       Rule.Calc ("Groups.times_class.times", eval_binop "#mult_"),
wneuper@59416
   684
	       Rule.Calc ("Atools.pow", eval_binop "#power_")
neuper@37954
   685
	       ],
wneuper@59416
   686
      scr = Rule.EmptyScr(*Rule.Prog ((Thm.term_of o the o (parse thy)) 
neuper@37954
   687
      scr_make_polytest)*)
wneuper@59406
   688
      }; 
neuper@37954
   689
neuper@37954
   690
val expand_binomtest =
wneuper@59416
   691
  Rule.Rls{id = "expand_binomtest", preconds = [], 
neuper@37954
   692
      rew_ord = ("termlessI",termlessI),
wneuper@59416
   693
      erls = testerls, srls = Rule.Erls,
neuper@38014
   694
      calc = [("PLUS"  , ("Groups.plus_class.plus", eval_binop "#add_")), 
neuper@38034
   695
	      ("TIMES" , ("Groups.times_class.times", eval_binop "#mult_")),
neuper@37954
   696
	      ("POWER", ("Atools.pow", eval_binop "#power_"))
neuper@42451
   697
	      ], errpatts = [],
neuper@38001
   698
      rules = 
wneuper@59416
   699
      [Rule.Thm ("real_plus_binom_pow2"  ,TermC.num_str @{thm real_plus_binom_pow2}),     
neuper@37954
   700
	       (*"(a + b) ^^^ 2 = a ^^^ 2 + 2 * a * b + b ^^^ 2"*)
wneuper@59416
   701
       Rule.Thm ("real_plus_binom_times" ,TermC.num_str @{thm real_plus_binom_times}),    
neuper@37954
   702
	      (*"(a + b)*(a + b) = ...*)
wneuper@59416
   703
       Rule.Thm ("real_minus_binom_pow2" ,TermC.num_str @{thm real_minus_binom_pow2}),   
neuper@38001
   704
       (*"(a - b) ^^^ 2 = a ^^^ 2 - 2 * a * b + b ^^^ 2"*)
wneuper@59416
   705
       Rule.Thm ("real_minus_binom_times",TermC.num_str @{thm real_minus_binom_times}),   
neuper@38001
   706
       (*"(a - b)*(a - b) = ...*)
wneuper@59416
   707
       Rule.Thm ("real_plus_minus_binom1",TermC.num_str @{thm real_plus_minus_binom1}),   
neuper@38001
   708
        (*"(a + b) * (a - b) = a ^^^ 2 - b ^^^ 2"*)
wneuper@59416
   709
       Rule.Thm ("real_plus_minus_binom2",TermC.num_str @{thm real_plus_minus_binom2}),   
neuper@38001
   710
        (*"(a - b) * (a + b) = a ^^^ 2 - b ^^^ 2"*)
neuper@38001
   711
       (*RL 020915*)
wneuper@59416
   712
       Rule.Thm ("real_pp_binom_times",TermC.num_str @{thm real_pp_binom_times}), 
neuper@38001
   713
        (*(a + b)*(c + d) = a*c + a*d + b*c + b*d*)
wneuper@59416
   714
       Rule.Thm ("real_pm_binom_times",TermC.num_str @{thm real_pm_binom_times}), 
neuper@38001
   715
        (*(a + b)*(c - d) = a*c - a*d + b*c - b*d*)
wneuper@59416
   716
       Rule.Thm ("real_mp_binom_times",TermC.num_str @{thm real_mp_binom_times}), 
neuper@38001
   717
        (*(a - b)*(c p d) = a*c + a*d - b*c - b*d*)
wneuper@59416
   718
       Rule.Thm ("real_mm_binom_times",TermC.num_str @{thm real_mm_binom_times}), 
neuper@38001
   719
        (*(a - b)*(c p d) = a*c - a*d - b*c + b*d*)
wneuper@59416
   720
       Rule.Thm ("realpow_multI",TermC.num_str @{thm realpow_multI}),                
neuper@38001
   721
        (*(a*b)^^^n = a^^^n * b^^^n*)
wneuper@59416
   722
       Rule.Thm ("real_plus_binom_pow3",TermC.num_str @{thm real_plus_binom_pow3}),
neuper@38001
   723
        (* (a + b)^^^3 = a^^^3 + 3*a^^^2*b + 3*a*b^^^2 + b^^^3 *)
wneuper@59416
   724
       Rule.Thm ("real_minus_binom_pow3",TermC.num_str @{thm real_minus_binom_pow3}),
neuper@38001
   725
        (* (a - b)^^^3 = a^^^3 - 3*a^^^2*b + 3*a*b^^^2 - b^^^3 *)
neuper@37954
   726
neuper@37954
   727
wneuper@59416
   728
     (*  Rule.Thm ("distrib_right" ,TermC.num_str @{thm distrib_right}),	
neuper@38001
   729
	 (*"(z1.0 + z2.0) * w = z1.0 * w + z2.0 * w"*)
wneuper@59416
   730
	Rule.Thm ("distrib_left",TermC.num_str @{thm distrib_left}),	
neuper@38001
   731
	(*"w * (z1.0 + z2.0) = w * z1.0 + w * z2.0"*)
wneuper@59416
   732
	Rule.Thm ("left_diff_distrib" ,TermC.num_str @{thm left_diff_distrib}),	
neuper@38001
   733
	(*"(z1.0 - z2.0) * w = z1.0 * w - z2.0 * w"*)
wneuper@59416
   734
	Rule.Thm ("right_diff_distrib",TermC.num_str @{thm right_diff_distrib}),	
neuper@38001
   735
	(*"w * (z1.0 - z2.0) = w * z1.0 - w * z2.0"*)
neuper@38001
   736
	*)
neuper@38001
   737
	
wneuper@59416
   738
	Rule.Thm ("mult_1_left",TermC.num_str @{thm mult_1_left}),              
neuper@38001
   739
         (*"1 * z = z"*)
wneuper@59416
   740
	Rule.Thm ("mult_zero_left",TermC.num_str @{thm mult_zero_left}),              
neuper@38001
   741
         (*"0 * z = 0"*)
wneuper@59416
   742
	Rule.Thm ("add_0_left",TermC.num_str @{thm add_0_left}),
neuper@38001
   743
         (*"0 + z = z"*)
neuper@37954
   744
wneuper@59416
   745
	Rule.Calc ("Groups.plus_class.plus", eval_binop "#add_"), 
wneuper@59416
   746
	Rule.Calc ("Groups.times_class.times", eval_binop "#mult_"),
wneuper@59416
   747
	Rule.Calc ("Atools.pow", eval_binop "#power_"),
neuper@38001
   748
        (*	       
wneuper@59416
   749
	 Rule.Thm ("mult_commute",TermC.num_str @{thm mult_commute}),		
neuper@38001
   750
        (*AC-rewriting*)
wneuper@59416
   751
	Rule.Thm ("real_mult_left_commute",TermC.num_str @{thm real_mult_left_commute}),
wneuper@59416
   752
	Rule.Thm ("mult_assoc",TermC.num_str @{thm mult.assoc}),
wneuper@59416
   753
	Rule.Thm ("add_commute",TermC.num_str @{thm add_commute}),	
wneuper@59416
   754
	Rule.Thm ("add_left_commute",TermC.num_str @{thm add_left_commute}),
wneuper@59416
   755
	Rule.Thm ("add_assoc",TermC.num_str @{thm add_assoc}),
neuper@38001
   756
	*)
neuper@38001
   757
	
wneuper@59416
   758
	Rule.Thm ("sym_realpow_twoI",
wneuper@59389
   759
              TermC.num_str (@{thm realpow_twoI} RS @{thm sym})),
neuper@38001
   760
	(*"r1 * r1 = r1 ^^^ 2"*)
wneuper@59416
   761
	Rule.Thm ("realpow_plus_1",TermC.num_str @{thm realpow_plus_1}),			
neuper@38001
   762
	(*"r * r ^^^ n = r ^^^ (n + 1)"*)
wneuper@59416
   763
	(*Rule.Thm ("sym_real_mult_2",
wneuper@59389
   764
                TermC.num_str (@{thm real_mult_2} RS @{thm sym})),
neuper@38001
   765
	(*"z1 + z1 = 2 * z1"*)*)
wneuper@59416
   766
	Rule.Thm ("real_mult_2_assoc",TermC.num_str @{thm real_mult_2_assoc}),		
neuper@38001
   767
	(*"z1 + (z1 + k) = 2 * z1 + k"*)
neuper@37954
   768
wneuper@59416
   769
	Rule.Thm ("real_num_collect",TermC.num_str @{thm real_num_collect}), 
neuper@38001
   770
	(*"[| l is_const; m is_const |] ==> l * n + m * n = (l + m) * n"*)
wneuper@59416
   771
	Rule.Thm ("real_num_collect_assoc",TermC.num_str @{thm real_num_collect_assoc}),	
neuper@38001
   772
	(*"[| l is_const; m is_const |] ==>  l * n + (m * n + k) =  (l + m) * n + k"*)
wneuper@59416
   773
	Rule.Thm ("real_one_collect",TermC.num_str @{thm real_one_collect}),		
neuper@38001
   774
	(*"m is_const ==> n + m * n = (1 + m) * n"*)
wneuper@59416
   775
	Rule.Thm ("real_one_collect_assoc",TermC.num_str @{thm real_one_collect_assoc}), 
neuper@38001
   776
	(*"m is_const ==> k + (n + m * n) = k + (1 + m) * n"*)
neuper@37954
   777
wneuper@59416
   778
	Rule.Calc ("Groups.plus_class.plus", eval_binop "#add_"), 
wneuper@59416
   779
	Rule.Calc ("Groups.times_class.times", eval_binop "#mult_"),
wneuper@59416
   780
	Rule.Calc ("Atools.pow", eval_binop "#power_")
neuper@38001
   781
	],
wneuper@59416
   782
      scr = Rule.EmptyScr
wneuper@59186
   783
(*Script ((Thm.term_of o the o (parse thy)) scr_expand_binomtest)*)
wneuper@59406
   784
      };      
neuper@37954
   785
*}
neuper@52125
   786
setup {* KEStore_Elems.add_rlss 
s1210629013@55444
   787
  [("make_polytest", (Context.theory_name @{theory}, prep_rls' make_polytest)), 
s1210629013@55444
   788
  ("expand_binomtest", (Context.theory_name @{theory}, prep_rls' expand_binomtest))] *}
wneuper@59430
   789
setup {* KEStore_Elems.add_rlss 
wneuper@59430
   790
  [("norm_equation", (Context.theory_name @{theory}, prep_rls' norm_equation)), 
wneuper@59430
   791
  ("ac_plus_times", (Context.theory_name @{theory}, prep_rls' ac_plus_times)), 
wneuper@59430
   792
  ("rearrange_assoc", (Context.theory_name @{theory}, prep_rls' rearrange_assoc))] *}
wneuper@59430
   793
wneuper@59430
   794
section {* problems *}
wneuper@59430
   795
setup {* KEStore_Elems.add_pbts
wneuper@59430
   796
  [(Specify.prep_pbt thy "pbl_test_uni_plain2" [] Celem.e_pblID
wneuper@59430
   797
    (["plain_square","univariate","equation","test"],
wneuper@59430
   798
      [("#Given" ,["equality e_e","solveFor v_v"]),
wneuper@59430
   799
        ("#Where" ,["(matches (?a + ?b*v_v ^^^2 = 0) e_e) |" ^
wneuper@59430
   800
	        "(matches (     ?b*v_v ^^^2 = 0) e_e) |" ^
wneuper@59430
   801
	        "(matches (?a +    v_v ^^^2 = 0) e_e) |" ^
wneuper@59430
   802
	        "(matches (        v_v ^^^2 = 0) e_e)"]),
wneuper@59430
   803
        ("#Find"  ,["solutions v_v'i'"])],
wneuper@59430
   804
      assoc_rls' @{theory} "matches", 
wneuper@59430
   805
      SOME "solve (e_e::bool, v_v)", [["Test","solve_plain_square"]])),
wneuper@59430
   806
    (Specify.prep_pbt thy "pbl_test_uni_poly" [] Celem.e_pblID
wneuper@59430
   807
      (["polynomial","univariate","equation","test"],
wneuper@59430
   808
        [("#Given" ,["equality (v_v ^^^2 + p_p * v_v + q__q = 0)","solveFor v_v"]),
wneuper@59430
   809
          ("#Where" ,["HOL.False"]),
wneuper@59430
   810
          ("#Find"  ,["solutions v_v'i'"])],
wneuper@59430
   811
        Rule.e_rls, SOME "solve (e_e::bool, v_v)", [])),
wneuper@59430
   812
    (Specify.prep_pbt thy "pbl_test_uni_poly_deg2" [] Celem.e_pblID
wneuper@59430
   813
      (["degree_two","polynomial","univariate","equation","test"],
wneuper@59430
   814
        [("#Given" ,["equality (v_v ^^^2 + p_p * v_v + q__q = 0)","solveFor v_v"]),
wneuper@59430
   815
          ("#Find"  ,["solutions v_v'i'"])],
wneuper@59430
   816
        Rule.e_rls, SOME "solve (v_v ^^^2 + p_p * v_v + q__q = 0, v_v)", [])),
wneuper@59430
   817
    (Specify.prep_pbt thy "pbl_test_uni_poly_deg2_pq" [] Celem.e_pblID
wneuper@59430
   818
      (["pq_formula","degree_two","polynomial","univariate","equation","test"],
wneuper@59430
   819
        [("#Given" ,["equality (v_v ^^^2 + p_p * v_v + q__q = 0)","solveFor v_v"]),
wneuper@59430
   820
          ("#Find"  ,["solutions v_v'i'"])],
wneuper@59430
   821
        Rule.e_rls, SOME "solve (v_v ^^^2 + p_p * v_v + q__q = 0, v_v)", [])),
wneuper@59430
   822
    (Specify.prep_pbt thy "pbl_test_uni_poly_deg2_abc" [] Celem.e_pblID
wneuper@59430
   823
      (["abc_formula","degree_two","polynomial","univariate","equation","test"],
wneuper@59430
   824
        [("#Given" ,["equality (a_a * x ^^^2 + b_b * x + c_c = 0)","solveFor v_v"]),
wneuper@59430
   825
          ("#Find"  ,["solutions v_v'i'"])],
wneuper@59430
   826
        Rule.e_rls, SOME "solve (a_a * x ^^^2 + b_b * x + c_c = 0, v_v)", [])),
wneuper@59430
   827
    (Specify.prep_pbt thy "pbl_test_uni_root" [] Celem.e_pblID
wneuper@59430
   828
      (["squareroot","univariate","equation","test"],
wneuper@59430
   829
        [("#Given" ,["equality e_e","solveFor v_v"]),
wneuper@59430
   830
          ("#Where" ,["precond_rootpbl v_v"]),
wneuper@59430
   831
          ("#Find"  ,["solutions v_v'i'"])],
wneuper@59430
   832
        Rule.append_rls "contains_root" Rule.e_rls [Rule.Calc ("Test.contains'_root",
wneuper@59430
   833
            eval_contains_root "#contains_root_")], 
wneuper@59430
   834
        SOME "solve (e_e::bool, v_v)", [["Test","square_equation"]])),
wneuper@59430
   835
    (Specify.prep_pbt thy "pbl_test_uni_norm" [] Celem.e_pblID
wneuper@59430
   836
      (["normalise","univariate","equation","test"],
wneuper@59430
   837
        [("#Given" ,["equality e_e","solveFor v_v"]),
wneuper@59430
   838
          ("#Where" ,[]),
wneuper@59430
   839
          ("#Find"  ,["solutions v_v'i'"])],
wneuper@59430
   840
        Rule.e_rls, SOME "solve (e_e::bool, v_v)", [["Test","norm_univar_equation"]])),
wneuper@59430
   841
    (Specify.prep_pbt thy "pbl_test_uni_roottest" [] Celem.e_pblID
wneuper@59430
   842
      (["sqroot-test","univariate","equation","test"],
wneuper@59430
   843
        [("#Given" ,["equality e_e","solveFor v_v"]),
wneuper@59430
   844
          ("#Where" ,["precond_rootpbl v_v"]),
wneuper@59430
   845
          ("#Find"  ,["solutions v_v'i'"])],
wneuper@59430
   846
        Rule.e_rls, SOME "solve (e_e::bool, v_v)", [])),
wneuper@59430
   847
    (Specify.prep_pbt thy "pbl_test_intsimp" [] Celem.e_pblID
wneuper@59430
   848
      (["inttype","test"],
wneuper@59430
   849
        [("#Given" ,["intTestGiven t_t"]),
wneuper@59430
   850
          ("#Where" ,[]),
wneuper@59430
   851
          ("#Find"  ,["intTestFind s_s"])],
wneuper@59430
   852
      Rule.e_rls, NONE, [["Test","intsimp"]]))] *}
wneuper@59430
   853
wneuper@59430
   854
section {* methods *}
wneuper@59430
   855
wneuper@59430
   856
partial_function (tailrec) dummy1 :: "real \<Rightarrow> real"
wneuper@59430
   857
  where "dummy1 xxx = xxx"
wneuper@59430
   858
setup {* KEStore_Elems.add_mets
wneuper@59430
   859
  [Specify.prep_met @{theory "Diff"} "met_test" [] Celem.e_metID
wneuper@59430
   860
      (["Test"], [],
wneuper@59430
   861
        {rew_ord'="tless_true",rls'=Atools_erls,calc = [], srls = Rule.e_rls, prls=Rule.e_rls,
wneuper@59430
   862
          crls=Atools_erls, errpats = [], nrls = Rule.e_rls}, "empty_script"),
wneuper@59430
   863
    Specify.prep_met thy "met_test_solvelin" [] Celem.e_metID
wneuper@59430
   864
      (["Test","solve_linear"],
wneuper@59430
   865
        [("#Given" ,["equality e_e","solveFor v_v"]),
wneuper@59430
   866
          ("#Where" ,["matches (?a = ?b) e_e"]),
wneuper@59430
   867
          ("#Find"  ,["solutions v_v'i'"])],
wneuper@59430
   868
        {rew_ord' = "e_rew_ord", rls' = tval_rls, srls = Rule.e_rls,
wneuper@59430
   869
          prls = assoc_rls' @{theory} "matches", calc = [], crls = tval_rls, errpats = [],
wneuper@59430
   870
          nrls = Test_simplify},
wneuper@59430
   871
        "Script Solve_linear (e_e::bool) (v_v::real)=                     " ^
wneuper@59430
   872
          "(let e_e =                                                       " ^
wneuper@59430
   873
          "  Repeat                                                         " ^
wneuper@59430
   874
          "    (((Rewrite_Set_Inst [(bdv,v_v::real)] isolate_bdv False) @@  " ^
wneuper@59430
   875
          "      (Rewrite_Set Test_simplify False))) e_e" ^
wneuper@59430
   876
          " in [e_e::bool])")]
wneuper@59430
   877
*}
wneuper@59430
   878
partial_function (tailrec) dummy2 :: "real \<Rightarrow> real"
wneuper@59430
   879
  where "dummy2 xxx = xxx"
wneuper@59430
   880
setup {* KEStore_Elems.add_mets
wneuper@59430
   881
  [Specify.prep_met thy  "met_test_sqrt" [] Celem.e_metID
wneuper@59430
   882
      (*root-equation, version for tests before 8.01.01*)
wneuper@59430
   883
      (["Test","sqrt-equ-test"],
wneuper@59430
   884
        [("#Given" ,["equality e_e","solveFor v_v"]),
wneuper@59430
   885
          ("#Where" ,["contains_root (e_e::bool)"]),
wneuper@59430
   886
          ("#Find"  ,["solutions v_v'i'"])],
wneuper@59430
   887
        {rew_ord'="e_rew_ord",rls'=tval_rls,
wneuper@59430
   888
          srls = Rule.append_rls "srls_contains_root" Rule.e_rls
wneuper@59430
   889
              [Rule.Calc ("Test.contains'_root",eval_contains_root "")],
wneuper@59430
   890
          prls = Rule.append_rls "prls_contains_root" Rule.e_rls 
wneuper@59430
   891
              [Rule.Calc ("Test.contains'_root",eval_contains_root "")],
wneuper@59430
   892
          calc=[], crls=tval_rls, errpats = [], nrls = Rule.e_rls (*,asm_rls=[],
wneuper@59430
   893
          asm_thm=[("square_equation_left",""), ("square_equation_right","")]*)},
wneuper@59430
   894
        "Script Solve_root_equation (e_e::bool) (v_v::real) =  " ^
wneuper@59430
   895
          "(let e_e = " ^
wneuper@59430
   896
          "   ((While (contains_root e_e) Do" ^
wneuper@59430
   897
          "      ((Rewrite square_equation_left True) @@" ^
wneuper@59430
   898
          "       (Try (Rewrite_Set Test_simplify False)) @@" ^
wneuper@59430
   899
          "       (Try (Rewrite_Set rearrange_assoc False)) @@" ^
wneuper@59430
   900
          "       (Try (Rewrite_Set isolate_root False)) @@" ^
wneuper@59430
   901
          "       (Try (Rewrite_Set Test_simplify False)))) @@" ^
wneuper@59430
   902
          "    (Try (Rewrite_Set norm_equation False)) @@" ^
wneuper@59430
   903
          "    (Try (Rewrite_Set Test_simplify False)) @@" ^
wneuper@59430
   904
          "    (Rewrite_Set_Inst [(bdv,v_v::real)] isolate_bdv False) @@" ^
wneuper@59430
   905
          "    (Try (Rewrite_Set Test_simplify False)))" ^
wneuper@59430
   906
          "   e_e" ^
wneuper@59430
   907
          " in [e_e::bool])")]
wneuper@59430
   908
*}
wneuper@59430
   909
partial_function (tailrec) dummy3 :: "real \<Rightarrow> real"
wneuper@59430
   910
  where "dummy3 xxx = xxx"
wneuper@59430
   911
setup {* KEStore_Elems.add_mets
wneuper@59430
   912
  [Specify.prep_met thy  "met_test_sqrt2" [] Celem.e_metID
wneuper@59430
   913
      (*root-equation ... for test-*.sml until 8.01*)
wneuper@59430
   914
      (["Test","squ-equ-test2"],
wneuper@59430
   915
        [("#Given" ,["equality e_e","solveFor v_v"]),
wneuper@59430
   916
          ("#Find"  ,["solutions v_v'i'"])],
wneuper@59430
   917
        {rew_ord'="e_rew_ord",rls'=tval_rls,
wneuper@59430
   918
          srls = Rule.append_rls "srls_contains_root" Rule.e_rls 
wneuper@59430
   919
              [Rule.Calc ("Test.contains'_root",eval_contains_root"")],
wneuper@59430
   920
          prls=Rule.e_rls,calc=[], crls=tval_rls, errpats = [], nrls = Rule.e_rls(*,asm_rls=[],
wneuper@59430
   921
          asm_thm=[("square_equation_left",""), ("square_equation_right","")]*)},
wneuper@59430
   922
        "Script Solve_root_equation (e_e::bool) (v_v::real) =  " ^
wneuper@59430
   923
          "(let e_e = " ^
wneuper@59430
   924
          "   ((While (contains_root e_e) Do" ^
wneuper@59430
   925
          "      ((Rewrite square_equation_left True) @@" ^
wneuper@59430
   926
          "       (Try (Rewrite_Set Test_simplify False)) @@" ^
wneuper@59430
   927
          "       (Try (Rewrite_Set rearrange_assoc False)) @@" ^
wneuper@59430
   928
          "       (Try (Rewrite_Set isolate_root False)) @@" ^
wneuper@59430
   929
          "       (Try (Rewrite_Set Test_simplify False)))) @@" ^
wneuper@59430
   930
          "    (Try (Rewrite_Set norm_equation False)) @@" ^
wneuper@59430
   931
          "    (Try (Rewrite_Set Test_simplify False)) @@" ^
wneuper@59430
   932
          "    (Rewrite_Set_Inst [(bdv,v_v::real)] isolate_bdv False) @@" ^
wneuper@59430
   933
          "    (Try (Rewrite_Set Test_simplify False)))" ^
wneuper@59430
   934
          "   e_e;" ^
wneuper@59430
   935
          "  (L_L::bool list) = Tac subproblem_equation_dummy;          " ^
wneuper@59430
   936
          "  L_L = Tac solve_equation_dummy                             " ^
wneuper@59430
   937
          "  in Check_elementwise L_L {(v_v::real). Assumptions})")]
wneuper@59430
   938
*}
wneuper@59430
   939
partial_function (tailrec) dummy4 :: "real \<Rightarrow> real"
wneuper@59430
   940
  where "dummy4 xxx = xxx"
wneuper@59430
   941
setup {* KEStore_Elems.add_mets
wneuper@59430
   942
  [Specify.prep_met thy "met_test_squ_sub" [] Celem.e_metID
wneuper@59430
   943
      (*tests subproblem fixed linear*)
wneuper@59430
   944
      (["Test","squ-equ-test-subpbl1"],
wneuper@59430
   945
        [("#Given" ,["equality e_e","solveFor v_v"]),
wneuper@59430
   946
          ("#Where" ,["precond_rootmet v_v"]),
wneuper@59430
   947
          ("#Find"  ,["solutions v_v'i'"])],
wneuper@59430
   948
        {rew_ord' = "e_rew_ord", rls' = tval_rls, srls = Rule.e_rls,
wneuper@59430
   949
          prls = Rule.append_rls "prls_met_test_squ_sub" Rule.e_rls
wneuper@59430
   950
              [Rule.Calc ("Test.precond'_rootmet", eval_precond_rootmet "")],
wneuper@59430
   951
          calc=[], crls=tval_rls, errpats = [], nrls = Test_simplify},
wneuper@59430
   952
        "Script Solve_root_equation (e_e::bool) (v_v::real) =       " ^
wneuper@59430
   953
        " (let e_e = ((Try (Rewrite_Set norm_equation False)) @@    " ^
wneuper@59430
   954
        "            (Try (Rewrite_Set Test_simplify False))) e_e;  " ^
wneuper@59430
   955
        "     (L_L::bool list) =                                    " ^
wneuper@59430
   956
        "            (SubProblem (Test',                            " ^
wneuper@59430
   957
        "                         [LINEAR,univariate,equation,test]," ^
wneuper@59430
   958
        "                         [Test,solve_linear])              " ^
wneuper@59430
   959
        "                        [BOOL e_e, REAL v_v])              " ^
wneuper@59430
   960
        "  in Check_elementwise L_L {(v_v::real). Assumptions})     ")]
wneuper@59430
   961
*}
wneuper@59430
   962
partial_function (tailrec) dummy5 :: "real \<Rightarrow> real"
wneuper@59430
   963
  where "dummy5 xxx = xxx"
wneuper@59430
   964
setup {* KEStore_Elems.add_mets
wneuper@59430
   965
  [Specify.prep_met thy "met_test_squ_sub2" [] Celem.e_metID
wneuper@59430
   966
      (*tests subproblem fixed degree 2*)
wneuper@59430
   967
      (["Test","squ-equ-test-subpbl2"],
wneuper@59430
   968
        [("#Given" ,["equality e_e","solveFor v_v"]),
wneuper@59430
   969
          ("#Find"  ,["solutions v_v'i'"])],
wneuper@59430
   970
        {rew_ord'="e_rew_ord",rls'=tval_rls,srls=Rule.e_rls,prls=Rule.e_rls,calc=[], crls=tval_rls,
wneuper@59430
   971
          errpats = [], nrls = Rule.e_rls (*, asm_rls=[],asm_thm=[("square_equation_left",""),
wneuper@59430
   972
          ("square_equation_right","")]*)},
wneuper@59430
   973
        "Script Solve_root_equation (e_e::bool) (v_v::real) =  " ^
wneuper@59430
   974
        " (let e_e = Try (Rewrite_Set norm_equation False) e_e;              " ^
wneuper@59430
   975
        "(L_L::bool list) = (SubProblem (Test',[LINEAR,univariate,equation,test]," ^
wneuper@59430
   976
        "                    [Test,solve_by_pq_formula]) [BOOL e_e, REAL v_v])" ^
wneuper@59430
   977
        "in Check_elementwise L_L {(v_v::real). Assumptions})")]
wneuper@59430
   978
*}
wneuper@59430
   979
partial_function (tailrec) dummy7 :: "real \<Rightarrow> real"
wneuper@59430
   980
  where "dummy7 xxx = xxx"
wneuper@59430
   981
setup {* KEStore_Elems.add_mets
wneuper@59430
   982
  [Specify.prep_met thy  "met_test_eq1" [] Celem.e_metID
wneuper@59430
   983
      (*root-equation1:*)
wneuper@59430
   984
      (["Test","square_equation1"],
wneuper@59430
   985
        [("#Given" ,["equality e_e","solveFor v_v"]),
wneuper@59430
   986
          ("#Find"  ,["solutions v_v'i'"])],
wneuper@59430
   987
        {rew_ord'="e_rew_ord",rls'=tval_rls,
wneuper@59430
   988
          srls = Rule.append_rls "srls_contains_root" Rule.e_rls 
wneuper@59430
   989
            [Rule.Calc ("Test.contains'_root",eval_contains_root"")], prls=Rule.e_rls, calc=[], crls=tval_rls,
wneuper@59430
   990
              errpats = [], nrls = Rule.e_rls(*,asm_rls=[], asm_thm=[("square_equation_left",""),
wneuper@59430
   991
              ("square_equation_right","")]*)},
wneuper@59430
   992
        "Script Solve_root_equation (e_e::bool) (v_v::real) =  " ^
wneuper@59430
   993
          "(let e_e = " ^
wneuper@59430
   994
          "   ((While (contains_root e_e) Do" ^
wneuper@59430
   995
          "      ((Rewrite square_equation_left True) @@" ^
wneuper@59430
   996
          "       (Try (Rewrite_Set Test_simplify False)) @@" ^
wneuper@59430
   997
          "       (Try (Rewrite_Set rearrange_assoc False)) @@" ^
wneuper@59430
   998
          "       (Try (Rewrite_Set isolate_root False)) @@" ^
wneuper@59430
   999
          "       (Try (Rewrite_Set Test_simplify False)))) @@" ^
wneuper@59430
  1000
          "    (Try (Rewrite_Set norm_equation False)) @@" ^
wneuper@59430
  1001
          "    (Try (Rewrite_Set Test_simplify False)))" ^
wneuper@59430
  1002
          "   e_e;" ^
wneuper@59430
  1003
          "  (L_L::bool list) = (SubProblem (Test',[LINEAR,univariate,equation,test]," ^
wneuper@59430
  1004
          "                    [Test,solve_linear]) [BOOL e_e, REAL v_v])" ^
wneuper@59430
  1005
          "  in Check_elementwise L_L {(v_v::real). Assumptions})")]
wneuper@59430
  1006
*}
wneuper@59430
  1007
partial_function (tailrec) dummy8 :: "real \<Rightarrow> real"
wneuper@59430
  1008
  where "dummy8 xxx = xxx"
wneuper@59430
  1009
setup {* KEStore_Elems.add_mets
wneuper@59430
  1010
  [Specify.prep_met thy "met_test_squ2" [] Celem.e_metID
wneuper@59430
  1011
      (*root-equation2*)
wneuper@59430
  1012
        (["Test","square_equation2"],
wneuper@59430
  1013
          [("#Given" ,["equality e_e","solveFor v_v"]),
wneuper@59430
  1014
          ("#Find"  ,["solutions v_v'i'"])],
wneuper@59430
  1015
        {rew_ord'="e_rew_ord",rls'=tval_rls,
wneuper@59430
  1016
          srls = Rule.append_rls "srls_contains_root" Rule.e_rls 
wneuper@59430
  1017
              [Rule.Calc ("Test.contains'_root",eval_contains_root"")],
wneuper@59430
  1018
          prls=Rule.e_rls,calc=[], crls=tval_rls, errpats = [], nrls = Rule.e_rls(*,asm_rls=[],
wneuper@59430
  1019
          asm_thm=[("square_equation_left",""), ("square_equation_right","")]*)},
wneuper@59430
  1020
        "Script Solve_root_equation (e_e::bool) (v_v::real)  =  " ^
wneuper@59430
  1021
        "(let e_e = " ^
wneuper@59430
  1022
        "   ((While (contains_root e_e) Do" ^
wneuper@59430
  1023
        "      (((Rewrite square_equation_left True) Or " ^
wneuper@59430
  1024
        "        (Rewrite square_equation_right True)) @@" ^
wneuper@59430
  1025
        "       (Try (Rewrite_Set Test_simplify False)) @@" ^
wneuper@59430
  1026
        "       (Try (Rewrite_Set rearrange_assoc False)) @@" ^
wneuper@59430
  1027
        "       (Try (Rewrite_Set isolate_root False)) @@" ^
wneuper@59430
  1028
        "       (Try (Rewrite_Set Test_simplify False)))) @@" ^
wneuper@59430
  1029
        "    (Try (Rewrite_Set norm_equation False)) @@" ^
wneuper@59430
  1030
        "    (Try (Rewrite_Set Test_simplify False)))" ^
wneuper@59430
  1031
        "   e_e;" ^
wneuper@59430
  1032
        "  (L_L::bool list) = (SubProblem (Test',[plain_square,univariate,equation,test]," ^
wneuper@59430
  1033
        "                    [Test,solve_plain_square]) [BOOL e_e, REAL v_v])" ^
wneuper@59430
  1034
        "  in Check_elementwise L_L {(v_v::real). Assumptions})")]
wneuper@59430
  1035
*}
wneuper@59430
  1036
partial_function (tailrec) dummy9 :: "real \<Rightarrow> real"
wneuper@59430
  1037
  where "dummy9 xxx = xxx"
wneuper@59430
  1038
setup {* KEStore_Elems.add_mets
wneuper@59430
  1039
  [Specify.prep_met thy "met_test_squeq" [] Celem.e_metID
wneuper@59430
  1040
      (*root-equation*)
wneuper@59430
  1041
      (["Test","square_equation"],
wneuper@59430
  1042
        [("#Given" ,["equality e_e","solveFor v_v"]),
wneuper@59430
  1043
          ("#Find"  ,["solutions v_v'i'"])],
wneuper@59430
  1044
        {rew_ord'="e_rew_ord",rls'=tval_rls,
wneuper@59430
  1045
          srls = Rule.append_rls "srls_contains_root" Rule.e_rls 
wneuper@59430
  1046
              [Rule.Calc ("Test.contains'_root",eval_contains_root"")],
wneuper@59430
  1047
          prls=Rule.e_rls,calc=[], crls=tval_rls, errpats = [], nrls = Rule.e_rls (*,asm_rls=[],
wneuper@59430
  1048
          asm_thm=[("square_equation_left",""), ("square_equation_right","")]*)},
wneuper@59430
  1049
        "Script Solve_root_equation (e_e::bool) (v_v::real) =  " ^
wneuper@59430
  1050
          "(let e_e = " ^
wneuper@59430
  1051
          "   ((While (contains_root e_e) Do" ^
wneuper@59430
  1052
          "      (((Rewrite square_equation_left True) Or" ^
wneuper@59430
  1053
          "        (Rewrite square_equation_right True)) @@" ^
wneuper@59430
  1054
          "       (Try (Rewrite_Set Test_simplify False)) @@" ^
wneuper@59430
  1055
          "       (Try (Rewrite_Set rearrange_assoc False)) @@" ^
wneuper@59430
  1056
          "       (Try (Rewrite_Set isolate_root False)) @@" ^
wneuper@59430
  1057
          "       (Try (Rewrite_Set Test_simplify False)))) @@" ^
wneuper@59430
  1058
          "    (Try (Rewrite_Set norm_equation False)) @@" ^
wneuper@59430
  1059
          "    (Try (Rewrite_Set Test_simplify False)))" ^
wneuper@59430
  1060
          "   e_e;" ^
wneuper@59430
  1061
          "  (L_L::bool list) = (SubProblem (Test',[univariate,equation,test]," ^
wneuper@59430
  1062
          "                    [no_met]) [BOOL e_e, REAL v_v])" ^
wneuper@59430
  1063
          "  in Check_elementwise L_L {(v_v::real). Assumptions})")]
wneuper@59430
  1064
*}
wneuper@59430
  1065
partial_function (tailrec) dummy10 :: "real \<Rightarrow> real"
wneuper@59430
  1066
  where "dummy10 xxx = xxx"
wneuper@59430
  1067
setup {* KEStore_Elems.add_mets
wneuper@59430
  1068
  [Specify.prep_met thy "met_test_eq_plain" [] Celem.e_metID
wneuper@59430
  1069
      (*solve_plain_square*)
wneuper@59430
  1070
      (["Test","solve_plain_square"],
wneuper@59430
  1071
        [("#Given",["equality e_e","solveFor v_v"]),
wneuper@59430
  1072
          ("#Where" ,["(matches (?a + ?b*v_v ^^^2 = 0) e_e) |" ^
wneuper@59430
  1073
              "(matches (     ?b*v_v ^^^2 = 0) e_e) |" ^
wneuper@59430
  1074
              "(matches (?a +    v_v ^^^2 = 0) e_e) |" ^
wneuper@59430
  1075
              "(matches (        v_v ^^^2 = 0) e_e)"]), 
wneuper@59430
  1076
          ("#Find"  ,["solutions v_v'i'"])],
wneuper@59430
  1077
        {rew_ord'="e_rew_ord",rls'=tval_rls,calc=[],srls=Rule.e_rls,
wneuper@59430
  1078
          prls = assoc_rls' @{theory} "matches", crls=tval_rls, errpats = [], nrls = Rule.e_rls(*,
wneuper@59430
  1079
          asm_rls=[],asm_thm=[]*)},
wneuper@59430
  1080
        "Script Solve_plain_square (e_e::bool) (v_v::real) =           " ^
wneuper@59430
  1081
          " (let e_e = ((Try (Rewrite_Set isolate_bdv False)) @@         " ^
wneuper@59430
  1082
          "            (Try (Rewrite_Set Test_simplify False)) @@     " ^
wneuper@59430
  1083
          "            ((Rewrite square_equality_0 False) Or        " ^
wneuper@59430
  1084
          "             (Rewrite square_equality True)) @@            " ^
wneuper@59430
  1085
          "            (Try (Rewrite_Set tval_rls False))) e_e             " ^
wneuper@59430
  1086
          "  in ((Or_to_List e_e)::bool list))")]
wneuper@59430
  1087
*}
wneuper@59430
  1088
partial_function (tailrec) dummy11 :: "real \<Rightarrow> real"
wneuper@59430
  1089
  where "dummy11 xxx = xxx"
wneuper@59430
  1090
setup {* KEStore_Elems.add_mets
wneuper@59430
  1091
  [Specify.prep_met thy "met_test_norm_univ" [] Celem.e_metID
wneuper@59430
  1092
      (["Test","norm_univar_equation"],
wneuper@59430
  1093
        [("#Given",["equality e_e","solveFor v_v"]),
wneuper@59430
  1094
          ("#Where" ,[]), 
wneuper@59430
  1095
          ("#Find"  ,["solutions v_v'i'"])],
wneuper@59430
  1096
        {rew_ord'="e_rew_ord",rls'=tval_rls,srls = Rule.e_rls,prls=Rule.e_rls, calc=[], crls=tval_rls,
wneuper@59430
  1097
          errpats = [], nrls = Rule.e_rls},
wneuper@59430
  1098
        "Script Norm_univar_equation (e_e::bool) (v_v::real) =      " ^
wneuper@59430
  1099
          " (let e_e = ((Try (Rewrite rnorm_equation_add False)) @@   " ^
wneuper@59430
  1100
          "            (Try (Rewrite_Set Test_simplify False))) e_e   " ^
wneuper@59430
  1101
          "  in (SubProblem (Test',[univariate,equation,test],         " ^
wneuper@59430
  1102
          "                    [no_met]) [BOOL e_e, REAL v_v]))")]
wneuper@59430
  1103
*}
wneuper@59430
  1104
partial_function (tailrec) dummy12 :: "real \<Rightarrow> real"
wneuper@59430
  1105
  where "dummy12 xxx = xxx"
wneuper@59430
  1106
setup {* KEStore_Elems.add_mets
wneuper@59430
  1107
  [Specify.prep_met thy "met_test_intsimp" [] Celem.e_metID
wneuper@59430
  1108
      (["Test","intsimp"],
wneuper@59430
  1109
        [("#Given" ,["intTestGiven t_t"]),
wneuper@59430
  1110
          ("#Where" ,[]),
wneuper@59430
  1111
          ("#Find"  ,["intTestFind s_s"])],
wneuper@59430
  1112
        {rew_ord' = "e_rew_ord", rls' = tval_rls, srls = Rule.e_rls,  prls = Rule.e_rls, calc = [],
wneuper@59430
  1113
          crls = tval_rls, errpats = [], nrls = Test_simplify},
wneuper@59430
  1114
        "Script STest_simplify (t_t::int) =                  " ^
wneuper@59430
  1115
          "(Repeat                                                          " ^
wneuper@59430
  1116
          "    ((Try (Calculate PLUS)) @@  " ^
wneuper@59430
  1117
          "     (Try (Calculate TIMES))) t_t::int)")]
wneuper@59430
  1118
*}
wneuper@59430
  1119
wneuper@59430
  1120
section {* THIS WAS FOR TESTS ONLY, WILL BE REPLACED BY KEStore_Elems.get_calcs *}
wneuper@59430
  1121
ML {*
wneuper@59430
  1122
(* association list for calculate_, calculate
wneuper@59430
  1123
   "Groups.plus_class.plus" etc. not usable in scripts *)
wneuper@59430
  1124
val calclist =
wneuper@59430
  1125
    [
wneuper@59430
  1126
     (*as Tools.ML*)
wneuper@59430
  1127
     ("Vars"    ,("Tools.Vars"    ,eval_var "#Vars_")),
wneuper@59430
  1128
     ("matches",("Tools.matches",eval_matches "#matches_")),
wneuper@59430
  1129
     ("lhs"    ,("Tools.lhs"    ,eval_lhs "")),
wneuper@59430
  1130
     (*aus Atools.ML*)
wneuper@59430
  1131
     ("PLUS"    ,("Groups.plus_class.plus"        ,eval_binop "#add_")),
wneuper@59430
  1132
     ("TIMES"   ,("Groups.times_class.times"        ,eval_binop "#mult_")),
wneuper@59430
  1133
     ("DIVIDE" ,("Rings.divide_class.divide"  ,eval_cancel "#divide_e")),
wneuper@59430
  1134
     ("POWER"  ,("Atools.pow"  ,eval_binop "#power_")),
wneuper@59430
  1135
     ("is_const",("Atools.is'_const",eval_const "#is_const_")),
wneuper@59430
  1136
     ("le"      ,("Orderings.ord_class.less"        ,eval_equ "#less_")),
wneuper@59430
  1137
     ("leq"     ,("Orderings.ord_class.less_eq"       ,eval_equ "#less_equal_")),
wneuper@59430
  1138
     ("ident"   ,("Atools.ident",eval_ident "#ident_")),
wneuper@59430
  1139
     (*von hier (ehem.SqRoot*)
wneuper@59430
  1140
     ("sqrt"    ,("NthRoot.sqrt"   ,eval_sqrt "#sqrt_")),
wneuper@59430
  1141
     ("Test.contains_root",("contains'_root", eval_contains_root "#contains'_root")),
wneuper@59430
  1142
     ("Test.contains_root",("contains'_root",
wneuper@59430
  1143
                eval_contains_root"#contains_root_"))
wneuper@59430
  1144
     ];
wneuper@59430
  1145
*}
neuper@37906
  1146
neuper@37906
  1147
end