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