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