src/Tools/isac/Knowledge/Poly.thy
author wneuper <Walther.Neuper@jku.at>
Thu, 04 Aug 2022 12:48:37 +0200
changeset 60509 2e0b7ca391dc
parent 60506 145e45cd7a0f
child 60515 03e19793d81e
permissions -rw-r--r--
polish naming in Rewrite_Order
neuper@37906
     1
(* WN.020812: theorems in the Reals,
neuper@37906
     2
   necessary for special rule sets, in addition to Isabelle2002.
neuper@37906
     3
   !!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!
neuper@37906
     4
   !!! THIS IS THE _least_ NUMBER OF ADDITIONAL THEOREMS !!!
neuper@37906
     5
   !!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!!
walther@60242
     6
   xxxI contain \<up> instead of ^ in the respective theorem xxx in 2002
neuper@37906
     7
   changed by: Richard Lang 020912
neuper@37906
     8
*)
neuper@37906
     9
neuper@37950
    10
theory Poly imports Simplify begin
neuper@37906
    11
wneuper@59523
    12
subsection \<open>remark on term-structure of polynomials\<close>
wneuper@59523
    13
text \<open>
wneuper@59523
    14
WN190319:
wneuper@59523
    15
the code below reflects missing coordination between two authors:
wneuper@59523
    16
* ML: built the equation solver; simple rule-sets, programs; better predicates for specifications.
wneuper@59523
    17
* MG: built simplification of polynomials with AC rewriting by ML code
wneuper@59523
    18
wneuper@59523
    19
WN020919:
wneuper@59523
    20
*** there are 5 kinds of expanded normalforms ***
wneuper@59523
    21
wneuper@59523
    22
[1] 'complete polynomial' (Komplettes Polynom), univariate
wneuper@59523
    23
   a_0 + a_1.x^1 +...+ a_n.x^n   not (a_n = 0)
wneuper@59523
    24
	        not (a_n = 0), some a_i may be zero (DON'T disappear),
wneuper@59523
    25
                variables in monomials lexicographically ordered and complete,
wneuper@59523
    26
                x written as 1*x^1, ...
wneuper@59523
    27
[2] 'polynomial' (Polynom), univariate and multivariate
wneuper@59523
    28
   a_0 + a_1.x +...+ a_n.x^n   not (a_n = 0)
wneuper@59523
    29
   a_0 + a_1.x_1.x_2^n_12...x_m^n_1m +...+  a_n.x_1^n.x_2^n_n2...x_m^n_nm
wneuper@59523
    30
	        not (a_n = 0), some a_i may be zero (ie. monomials disappear),
wneuper@59523
    31
                exponents and coefficients equal 1 are not (WN060904.TODO in cancel_p_)shown,
wneuper@59523
    32
                and variables in monomials are lexicographically ordered  
walther@60242
    33
   examples: [1]: "1 + (-10) * x \<up> 1 + 25 * x \<up> 2"
walther@60242
    34
	     [1]: "11 + 0 * x \<up> 1 + 1 * x \<up> 2"
walther@60242
    35
	     [2]: "x + (-50) * x \<up> 3"
walther@60242
    36
	     [2]: "(-1) * x * y \<up> 2 + 7 * x \<up> 3"
wneuper@59523
    37
wneuper@59523
    38
[3] 'expanded_term' (Ausmultiplizierter Term):
wneuper@59523
    39
   pull out unary minus to binary minus, 
wneuper@59523
    40
   as frequently exercised in schools; other conditions for [2] hold however
walther@60242
    41
   examples: "a \<up> 2 - 2 * a * b + b \<up> 2"
walther@60242
    42
	     "4 * x \<up> 2 - 9 * y \<up> 2"
wneuper@59523
    43
[4] 'polynomial_in' (Polynom in): 
wneuper@59523
    44
   polynomial in 1 variable with arbitrary coefficients
walther@60242
    45
   examples: "2 * x + (-50) * x \<up> 3"                     (poly in x)
walther@60242
    46
	     "(u + v) + (2 * u \<up> 2) * a + (-u) * a \<up> 2 (poly in a)
wneuper@59523
    47
[5] 'expanded_in' (Ausmultiplizierter Termin in): 
wneuper@59523
    48
   analoguous to [3] with binary minus like [3]
walther@60242
    49
   examples: "2 * x - 50 * x \<up> 3"                     (expanded in x)
walther@60242
    50
	     "(u + v) + (2 * u \<up> 2) * a - u * a \<up> 2 (expanded in a)
wneuper@59523
    51
\<close>
wneuper@59523
    52
subsection \<open>consts definition for predicates in specifications\<close>
neuper@37906
    53
consts
neuper@37906
    54
walther@60278
    55
  is_expanded_in :: "[real, real] => bool" ("_ is'_expanded'_in _") 
walther@60278
    56
  is_poly_in     :: "[real, real] => bool" ("_ is'_poly'_in _")   (*RL DA *)
walther@60278
    57
  has_degree_in  :: "[real, real] => real" ("_ has'_degree'_in _")(*RL DA *)
walther@60278
    58
  is_polyrat_in  :: "[real, real] => bool" ("_ is'_polyrat'_in _")(*RL030626*)
neuper@37906
    59
walther@60278
    60
  is_multUnordered:: "real => bool" ("_ is'_multUnordered") 
walther@60278
    61
  is_addUnordered :: "real => bool" ("_ is'_addUnordered") (*WN030618*)
walther@60278
    62
  is_polyexp      :: "real => bool" ("_ is'_polyexp") 
neuper@37906
    63
wneuper@59523
    64
subsection \<open>theorems not yet adopted from Isabelle\<close>
neuper@52148
    65
axiomatization where (*.not contained in Isabelle2002,
neuper@37906
    66
         stated as axioms, TODO: prove as theorems;
walther@60242
    67
         theorem-IDs 'xxxI' with \<up> instead of ^ in 'xxx' in Isabelle2002.*)
neuper@37906
    68
walther@60242
    69
  realpow_pow:             "(a \<up> b) \<up> c = a \<up> (b * c)" and
walther@60242
    70
  realpow_addI:            "r \<up> (n + m) = r \<up> n * r \<up> m" and
walther@60242
    71
  realpow_addI_assoc_l:    "r \<up> n * (r \<up> m * s) = r \<up> (n + m) * s" and
walther@60242
    72
  realpow_addI_assoc_r:    "s * r \<up> n * r \<up> m = s * r \<up> (n + m)" and
neuper@37906
    73
		  
walther@60242
    74
  realpow_oneI:            "r \<up> 1 = r" and
walther@60242
    75
  realpow_zeroI:            "r \<up> 0 = 1" and
walther@60242
    76
  realpow_eq_oneI:         "1 \<up> n = 1" and
walther@60242
    77
  realpow_multI:           "(r * s) \<up> n = r \<up> n * s \<up> n"  and
neuper@37974
    78
  realpow_multI_poly:      "[| r is_polyexp; s is_polyexp |] ==>
walther@60242
    79
			      (r * s) \<up> n = r \<up> n * s \<up> n"  and
walther@60242
    80
  realpow_minus_oneI:      "(- 1) \<up> (2 * n) = 1"  and 
wneuper@59589
    81
  real_diff_0:		         "0 - x = - (x::real)" and
neuper@37906
    82
walther@60242
    83
  realpow_twoI:            "r \<up> 2 = r * r" and
walther@60242
    84
  realpow_twoI_assoc_l:	  "r * (r * s) = r \<up> 2 * s" and
walther@60242
    85
  realpow_twoI_assoc_r:	  "s * r * r = s * r \<up> 2" and
walther@60242
    86
  realpow_two_atom:        "r is_atom ==> r * r = r \<up> 2" and
walther@60242
    87
  realpow_plus_1:          "r * r \<up> n = r \<up> (n + 1)"   and       
walther@60242
    88
  realpow_plus_1_assoc_l:  "r * (r \<up> m * s) = r \<up> (1 + m) * s"  and
walther@60242
    89
  realpow_plus_1_assoc_l2: "r \<up> m * (r * s) = r \<up> (1 + m) * s"  and
walther@60242
    90
  realpow_plus_1_assoc_r:  "s * r * r \<up> m = s * r \<up> (1 + m)" and
walther@60242
    91
  realpow_plus_1_atom:     "r is_atom ==> r * r \<up> n = r \<up> (1 + n)" and
neuper@37974
    92
  realpow_def_atom:        "[| Not (r is_atom); 1 < n |]
walther@60242
    93
			   ==> r \<up> n = r * r \<up> (n + -1)" and
walther@60242
    94
  realpow_addI_atom:       "r is_atom ==> r \<up> n * r \<up> m = r \<up> (n + m)" and
neuper@37906
    95
neuper@37906
    96
walther@60278
    97
  realpow_minus_even:	     "n is_even ==> (- r) \<up> n = r \<up> n" and
walther@60242
    98
  realpow_minus_odd:       "Not (n is_even) ==> (- r) \<up> n = -1 * r \<up> n" and
neuper@37906
    99
neuper@37906
   100
neuper@37906
   101
(* RL 020914 *)
neuper@52148
   102
  real_pp_binom_times:     "(a + b)*(c + d) = a*c + a*d + b*c + b*d" and
neuper@52148
   103
  real_pm_binom_times:     "(a + b)*(c - d) = a*c - a*d + b*c - b*d" and
neuper@52148
   104
  real_mp_binom_times:     "(a - b)*(c + d) = a*c + a*d - b*c - b*d" and
neuper@52148
   105
  real_mm_binom_times:     "(a - b)*(c - d) = a*c - a*d - b*c + b*d" and
walther@60242
   106
  real_plus_binom_pow3:    "(a + b) \<up> 3 = a \<up> 3 + 3*a \<up> 2*b + 3*a*b \<up> 2 + b \<up> 3" and
neuper@37974
   107
  real_plus_binom_pow3_poly: "[| a is_polyexp; b is_polyexp |] ==> 
walther@60242
   108
			    (a + b) \<up> 3 = a \<up> 3 + 3*a \<up> 2*b + 3*a*b \<up> 2 + b \<up> 3" and
walther@60242
   109
  real_minus_binom_pow3:   "(a - b) \<up> 3 = a \<up> 3 - 3*a \<up> 2*b + 3*a*b \<up> 2 - b \<up> 3" and
walther@60242
   110
  real_minus_binom_pow3_p: "(a + -1 * b) \<up> 3 = a \<up> 3 + -3*a \<up> 2*b + 3*a*b \<up> 2 +
walther@60242
   111
                           -1*b \<up> 3" and
walther@60385
   112
(* real_plus_binom_pow:        "[| n is_num;  3 < n |] ==>
walther@60260
   113
			       (a + b) \<up> n = (a + b) * (a + b)\<up>(n - 1)" *)
walther@60242
   114
  real_plus_binom_pow4:   "(a + b) \<up> 4 = (a \<up> 3 + 3*a \<up> 2*b + 3*a*b \<up> 2 + b \<up> 3)
neuper@52148
   115
                           *(a + b)" and
neuper@37974
   116
  real_plus_binom_pow4_poly: "[| a is_polyexp; b is_polyexp |] ==> 
walther@60242
   117
			   (a + b) \<up> 4 = (a \<up> 3 + 3*a \<up> 2*b + 3*a*b \<up> 2 + b \<up> 3)
neuper@52148
   118
                           *(a + b)" and
walther@60242
   119
  real_plus_binom_pow5:    "(a + b) \<up> 5 = (a \<up> 3 + 3*a \<up> 2*b + 3*a*b \<up> 2 + b \<up> 3)
walther@60242
   120
                           *(a \<up> 2 + 2*a*b + b \<up> 2)" and
neuper@37974
   121
  real_plus_binom_pow5_poly: "[| a is_polyexp; b is_polyexp |] ==> 
walther@60242
   122
			        (a + b) \<up> 5 = (a \<up> 3 + 3*a \<up> 2*b + 3*a*b \<up> 2 
walther@60242
   123
                                + b \<up> 3)*(a \<up> 2 + 2*a*b + b \<up> 2)" and
neuper@52148
   124
  real_diff_plus:          "a - b = a + -b" (*17.3.03: do_NOT_use*) and
neuper@52148
   125
  real_diff_minus:         "a - b = a + -1 * b" and
walther@60242
   126
  real_plus_binom_times:   "(a + b)*(a + b) = a \<up> 2 + 2*a*b + b \<up> 2" and
walther@60242
   127
  real_minus_binom_times:  "(a - b)*(a - b) = a \<up> 2 - 2*a*b + b \<up> 2" and
neuper@37906
   128
  (*WN071229 changed for Schaerding -----vvv*)
walther@60260
   129
  (*real_plus_binom_pow2:  "(a + b) \<up> 2 = a \<up> 2 + 2*a*b + b \<up> 2"*)
walther@60242
   130
  real_plus_binom_pow2:    "(a + b) \<up> 2 = (a + b) * (a + b)" and
walther@60242
   131
  (*WN071229 changed for Schaerding -----\<up>*)
neuper@37974
   132
  real_plus_binom_pow2_poly: "[| a is_polyexp; b is_polyexp |] ==>
walther@60242
   133
			       (a + b) \<up> 2 = a \<up> 2 + 2*a*b + b \<up> 2" and
walther@60242
   134
  real_minus_binom_pow2:      "(a - b) \<up> 2 = a \<up> 2 - 2*a*b + b \<up> 2" and
walther@60242
   135
  real_minus_binom_pow2_p:    "(a - b) \<up> 2 = a \<up> 2 + -2*a*b + b \<up> 2" and
walther@60242
   136
  real_plus_minus_binom1:     "(a + b)*(a - b) = a \<up> 2 - b \<up> 2" and
walther@60242
   137
  real_plus_minus_binom1_p:   "(a + b)*(a - b) = a \<up> 2 + -1*b \<up> 2" and
walther@60242
   138
  real_plus_minus_binom1_p_p: "(a + b)*(a + -1 * b) = a \<up> 2 + -1*b \<up> 2" and
walther@60242
   139
  real_plus_minus_binom2:     "(a - b)*(a + b) = a \<up> 2 - b \<up> 2" and
walther@60242
   140
  real_plus_minus_binom2_p:   "(a - b)*(a + b) = a \<up> 2 + -1*b \<up> 2" and
walther@60242
   141
  real_plus_minus_binom2_p_p: "(a + -1 * b)*(a + b) = a \<up> 2 + -1*b \<up> 2" and
walther@60242
   142
  real_plus_binom_times1:     "(a +  1*b)*(a + -1*b) = a \<up> 2 + -1*b \<up> 2" and
walther@60242
   143
  real_plus_binom_times2:     "(a + -1*b)*(a +  1*b) = a \<up> 2 + -1*b \<up> 2" and
neuper@37906
   144
walther@60385
   145
  real_num_collect:           "[| l is_num; m is_num |] ==>
neuper@52148
   146
			      l * n + m * n = (l + m) * n" and
neuper@37906
   147
(* FIXME.MG.0401: replace 'real_num_collect_assoc' 
neuper@37906
   148
	by 'real_num_collect_assoc_l' ... are equal, introduced by MG ! *)
walther@60385
   149
  real_num_collect_assoc:     "[| l is_num; m is_num |] ==> 
neuper@52148
   150
			      l * n + (m * n + k) = (l + m) * n + k" and
walther@60385
   151
  real_num_collect_assoc_l:   "[| l is_num; m is_num |] ==>
neuper@37950
   152
			      l * n + (m * n + k) = (l + m)
neuper@52148
   153
				* n + k" and
walther@60385
   154
  real_num_collect_assoc_r:   "[| l is_num; m is_num |] ==>
neuper@52148
   155
			      (k + m * n) + l * n = k + (l + m) * n" and
walther@60385
   156
  real_one_collect:           "m is_num ==> n + m * n = (1 + m) * n" and
neuper@37906
   157
(* FIXME.MG.0401: replace 'real_one_collect_assoc' 
neuper@37906
   158
	by 'real_one_collect_assoc_l' ... are equal, introduced by MG ! *)
walther@60385
   159
  real_one_collect_assoc:     "m is_num ==> n + (m * n + k) = (1 + m)* n + k" and
neuper@37906
   160
walther@60385
   161
  real_one_collect_assoc_l:   "m is_num ==> n + (m * n + k) = (1 + m) * n + k" and
walther@60385
   162
  real_one_collect_assoc_r:  "m is_num ==> (k + n) +  m * n = k + (1 + m) * n" and
neuper@37906
   163
neuper@37906
   164
(* FIXME.MG.0401: replace 'real_mult_2_assoc' 
neuper@37906
   165
	by 'real_mult_2_assoc_l' ... are equal, introduced by MG ! *)
neuper@52148
   166
  real_mult_2_assoc:          "z1 + (z1 + k) = 2 * z1 + k" and
neuper@52148
   167
  real_mult_2_assoc_l:        "z1 + (z1 + k) = 2 * z1 + k" and
neuper@52148
   168
  real_mult_2_assoc_r:        "(k + z1) + z1 = k + 2 * z1" and
neuper@37906
   169
wneuper@59587
   170
  real_mult_left_commute: "z1 * (z2 * z3) = z2 * (z1 * z3)" and
wneuper@59587
   171
  real_mult_minus1:       "-1 * z = - (z::real)" and
walther@60333
   172
  (*sym_real_mult_minus1 expands indefinitely without assumptions ...*)
walther@60343
   173
  real_mult_minus1_sym:   "[| \<not>(matches (- 1 * x) z); \<not>(z is_num) |] ==> - (z::real) = -1 * z" and
walther@60344
   174
  real_minus_mult_left:   "\<not> ((a::real) is_num) ==> (- a) * b = - (a * b)" and
wneuper@59587
   175
  real_mult_2:            "2 * z = z + (z::real)" and
walther@60343
   176
                                                                                                            
neuper@52148
   177
  real_add_mult_distrib_poly: "w is_polyexp ==> (z1 + z2) * w = z1 * w + z2 * w" and
neuper@37974
   178
  real_add_mult_distrib2_poly:"w is_polyexp ==> w * (z1 + z2) = w * z1 + w * z2"
neuper@37950
   179
walther@60278
   180
wneuper@59523
   181
subsection \<open>auxiliary functions\<close>
wneuper@59530
   182
ML \<open>
wneuper@59530
   183
val poly_consts =
wenzelm@60309
   184
  [\<^const_name>\<open>plus\<close>, \<^const_name>\<open>minus\<close>,
wenzelm@60309
   185
  \<^const_name>\<open>divide\<close>, \<^const_name>\<open>times\<close>,
wenzelm@60405
   186
  \<^const_name>\<open>realpow\<close>];
walther@60321
   187
walther@60321
   188
val int_ord_SAVE = int_ord;
walther@60321
   189
(*for tests on rewrite orders*)
walther@60321
   190
fun int_ord (i1, i2) =
walther@60321
   191
(@{print} {a = "int_ord (" ^ string_of_int i1 ^ ", " ^ string_of_int i2 ^ ") = ", z = Int.compare (i1, i2)};
walther@60321
   192
  Int.compare (i1, i2));
walther@60321
   193
(**)val int_ord = int_ord_SAVE; (*..outcomment for tests*)
wneuper@59530
   194
\<close>
wneuper@59523
   195
subsubsection \<open>for predicates in specifications (ML)\<close>
wneuper@59472
   196
ML \<open>
wneuper@59522
   197
(*--- auxiliary for is_expanded_in, is_poly_in, has_degree_in ---*)
walther@60317
   198
(*. a "monomial t in variable v" is a term t with
wneuper@59522
   199
  either (1) v NOT existent in t, or (2) v contained in t,
wneuper@59522
   200
  if (1) then degree 0
walther@60317
   201
  if (2) then v is a factor on the very right, casually with exponent.*)
wneuper@59522
   202
fun factor_right_deg (*case 2*)
walther@60335
   203
	    (Const (\<^const_name>\<open>Groups.times_class.times\<close>, _) $
wenzelm@60405
   204
        t1 $ (Const (\<^const_name>\<open>realpow\<close>,_) $ vv $ num)) v =
walther@60317
   205
	  if vv = v andalso not (Prog_Expr.occurs_in v t1) then SOME (snd (HOLogic.dest_number num))
walther@60317
   206
    else NONE
wenzelm@60405
   207
  | factor_right_deg (Const (\<^const_name>\<open>realpow\<close>,_) $ vv $ num) v =
walther@60317
   208
	   if (vv = v) then SOME (snd (HOLogic.dest_number num)) else NONE
walther@60331
   209
wenzelm@60309
   210
  | factor_right_deg (Const (\<^const_name>\<open>times\<close>,_) $ t1 $ vv) v = 
walther@59603
   211
	   if vv = v andalso not (Prog_Expr.occurs_in v t1) then SOME 1 else NONE
wneuper@59522
   212
  | factor_right_deg vv v =
wneuper@59522
   213
	  if (vv = v) then SOME 1 else NONE;    
wneuper@59522
   214
fun mono_deg_in m v =  (*case 1*)
walther@59603
   215
	if not (Prog_Expr.occurs_in v m) then (*case 1*) SOME 0 else factor_right_deg m v;
wneuper@59522
   216
wneuper@59522
   217
fun expand_deg_in t v =
wneuper@59522
   218
	let
wenzelm@60309
   219
    fun edi ~1 ~1 (Const (\<^const_name>\<open>plus\<close>, _) $ t1 $ t2) =
wneuper@59522
   220
          (case mono_deg_in t2 v of (* $ is left associative*)
wneuper@59522
   221
            SOME d' => edi d' d' t1 | NONE => NONE)
wenzelm@60309
   222
      | edi ~1 ~1 (Const (\<^const_name>\<open>minus\<close>, _) $ t1 $ t2) =
wneuper@59522
   223
          (case mono_deg_in t2 v of
wneuper@59522
   224
            SOME d' => edi d' d' t1 | NONE => NONE)
wenzelm@60309
   225
      | edi d dmax (Const (\<^const_name>\<open>minus\<close>, _) $ t1 $ t2) =
wneuper@59522
   226
          (case mono_deg_in t2 v of (*(d = 0 andalso d' = 0) handle 3+4-...4 +x*)
wneuper@59522
   227
	        SOME d' => if d > d' orelse (d = 0 andalso d' = 0) then edi d' dmax t1 else NONE
wneuper@59522
   228
          | NONE => NONE)
wenzelm@60309
   229
      | edi d dmax (Const (\<^const_name>\<open>plus\<close>,_) $ t1 $ t2) =
wneuper@59522
   230
          (case mono_deg_in t2 v of
wneuper@59522
   231
            SOME d' =>    (*RL (d = 0 andalso d' = 0) need to handle 3+4-...4 +x*)
wneuper@59522
   232
              if d > d' orelse (d = 0 andalso d' = 0) then edi d' dmax t1 else NONE
wneuper@59522
   233
          | NONE => NONE)
wneuper@59522
   234
      | edi ~1 ~1 t =
wneuper@59522
   235
          (case mono_deg_in t v of d as SOME _ => d | NONE => NONE)
wneuper@59522
   236
      | edi d dmax t = (*basecase last*)
wneuper@59522
   237
    	    (case mono_deg_in t v of
wneuper@59522
   238
    	      SOME d' => if d > d' orelse (d = 0 andalso d' = 0) then SOME dmax else NONE
wneuper@59522
   239
		      | NONE => NONE)
wneuper@59522
   240
	in edi ~1 ~1 t end;
wneuper@59522
   241
wneuper@59522
   242
fun poly_deg_in t v =
wneuper@59522
   243
	let
wenzelm@60309
   244
    fun edi ~1 ~1 (Const (\<^const_name>\<open>plus\<close>,_) $ t1 $ t2) =
wneuper@59522
   245
		    (case mono_deg_in t2 v of (* $ is left associative *)
wneuper@59522
   246
		      SOME d' => edi d' d' t1
wneuper@59522
   247
        | NONE => NONE)
wenzelm@60309
   248
	    | edi d dmax (Const (\<^const_name>\<open>plus\<close>,_) $ t1 $ t2) =
wneuper@59522
   249
		    (case mono_deg_in t2 v of
wneuper@59522
   250
	        SOME d' =>    (*RL (d = 0 andalso (d' = 0)) handle 3+4-...4 +x*)
wneuper@59522
   251
            if d > d' orelse (d = 0 andalso d' = 0) then edi d' dmax t1 else NONE
wneuper@59522
   252
        | NONE => NONE)
wneuper@59522
   253
	    | edi ~1 ~1 t =
wneuper@59522
   254
        (case mono_deg_in t v of
wneuper@59522
   255
		      d as SOME _ => d
wneuper@59522
   256
        | NONE => NONE)
wneuper@59522
   257
	    | edi d dmax t = (* basecase last *)
wneuper@59522
   258
		    (case mono_deg_in t v of
wneuper@59522
   259
		      SOME d' =>
wneuper@59522
   260
            if d > d' orelse (d = 0 andalso d' = 0) then SOME dmax else NONE
wneuper@59522
   261
        | NONE => NONE)
wneuper@59522
   262
	in edi ~1 ~1 t end;
wneuper@59523
   263
\<close>
neuper@37950
   264
wneuper@59523
   265
subsubsection \<open>for hard-coded AC rewriting (MG)\<close>
wneuper@59523
   266
ML \<open>
wneuper@59523
   267
(**. MG.03: make_polynomial_ ... uses SML-fun for ordering .**)
neuper@37950
   268
wneuper@59523
   269
(*FIXME.0401: make SML-order local to make_polynomial(_) *)
wneuper@59523
   270
(*FIXME.0401: replace 'make_polynomial'(old) by 'make_polynomial_'(MG) *)
wneuper@59523
   271
(* Polynom --> List von Monomen *) 
wenzelm@60309
   272
fun poly2list (Const (\<^const_name>\<open>plus\<close>,_) $ t1 $ t2) = 
wneuper@59523
   273
    (poly2list t1) @ (poly2list t2)
wneuper@59523
   274
  | poly2list t = [t];
neuper@37950
   275
wneuper@59523
   276
(* Monom --> Liste von Variablen *)
wenzelm@60309
   277
fun monom2list (Const (\<^const_name>\<open>times\<close>,_) $ t1 $ t2) = 
wneuper@59523
   278
    (monom2list t1) @ (monom2list t2)
wneuper@59523
   279
  | monom2list t = [t];
neuper@37950
   280
wneuper@59523
   281
(* liefert Variablenname (String) einer Variablen und Basis bei Potenz *)
wenzelm@60405
   282
fun get_basStr (Const (\<^const_name>\<open>realpow\<close>,_) $ Free (str, _) $ _) = str
wenzelm@60405
   283
  | get_basStr (Const (\<^const_name>\<open>realpow\<close>,_) $ n $ _) = TermC.to_string n
wneuper@59523
   284
  | get_basStr (Free (str, _)) = str
walther@60318
   285
  | get_basStr t =
walther@60322
   286
    if TermC.is_num t then TermC.to_string t
walther@60318
   287
    else "|||"; (* gross gewichtet; für Brüche ect. *)
neuper@37950
   288
wneuper@59523
   289
(* liefert Hochzahl (String) einer Variablen bzw Gewichtstring (zum Sortieren) *)
wenzelm@60405
   290
fun get_potStr (Const (\<^const_name>\<open>realpow\<close>, _) $ Free _ $ Free (str, _)) = str
wenzelm@60405
   291
  | get_potStr (Const (\<^const_name>\<open>realpow\<close>, _) $ Free _ $ t) =
walther@60322
   292
    if TermC.is_num t then TermC.to_string t else "|||"
walther@60321
   293
  | get_potStr (Free _) = "---" (* keine Hochzahl --> kleinst gewichtet *)
wneuper@59523
   294
  | get_potStr _ = "||||||"; (* gross gewichtet; für Brüch ect. *)
neuper@37978
   295
wneuper@59523
   296
(* Umgekehrte string_ord *)
wneuper@59523
   297
val string_ord_rev =  rev_order o string_ord;
wneuper@59523
   298
		
wneuper@59523
   299
 (* Ordnung zum lexikographischen Vergleich zweier Variablen (oder Potenzen) 
wneuper@59523
   300
    innerhalb eines Monomes:
wneuper@59523
   301
    - zuerst lexikographisch nach Variablenname 
wneuper@59523
   302
    - wenn gleich: nach steigender Potenz *)
walther@60321
   303
fun var_ord (a, b) = 
walther@60321
   304
(@{print} {a = "var_ord ", a_b = "(" ^ UnparseC.term a ^ ", " ^ UnparseC.term b ^ ")",
walther@60321
   305
   sort_args = "(" ^ get_basStr a ^ ", " ^ get_potStr a ^ "), (" ^ get_basStr b ^ ", " ^ get_potStr b ^ ")"};
walther@60321
   306
  prod_ord string_ord string_ord 
walther@60321
   307
  ((get_basStr a, get_potStr a), (get_basStr b, get_potStr b))
walther@60321
   308
);
walther@60318
   309
fun var_ord (a,b: term) = 
walther@60318
   310
  prod_ord string_ord string_ord 
wneuper@59523
   311
    ((get_basStr a, get_potStr a), (get_basStr b, get_potStr b));
neuper@37950
   312
wneuper@59523
   313
(* Ordnung zum lexikographischen Vergleich zweier Variablen (oder Potenzen); 
wneuper@59523
   314
   verwendet zum Sortieren von Monomen mittels Gesamtgradordnung:
wneuper@59523
   315
   - zuerst lexikographisch nach Variablenname 
wneuper@59523
   316
   - wenn gleich: nach sinkender Potenz*)
walther@60321
   317
fun var_ord_revPow (a, b: term) = 
walther@60321
   318
(@{print} {a = "var_ord_revPow ", at_bt = "(" ^ UnparseC.term a ^ ", " ^ UnparseC.term b ^ ")",
walther@60321
   319
    sort_args = "(" ^ get_basStr a ^ ", " ^ get_potStr a ^ "), (" ^ get_basStr b ^ ", " ^ get_potStr b ^ ")"};
walther@60321
   320
  prod_ord string_ord string_ord_rev 
walther@60321
   321
    ((get_basStr a, get_potStr a), (get_basStr b, get_potStr b))
walther@60321
   322
);
walther@60318
   323
fun var_ord_revPow (a, b: term) =
walther@60318
   324
  prod_ord string_ord string_ord_rev 
wneuper@59523
   325
    ((get_basStr a, get_potStr a), (get_basStr b, get_potStr b));
neuper@37950
   326
walther@60278
   327
wneuper@59523
   328
(* Ordnet ein Liste von Variablen (und Potenzen) lexikographisch *)
walther@60321
   329
fun sort_varList ts =
walther@60321
   330
(@{print} {a = "sort_varList", args = UnparseC.terms ts};
walther@60321
   331
  sort var_ord ts);
wneuper@59523
   332
val sort_varList = sort var_ord;
wneuper@59523
   333
wneuper@59523
   334
(* Entfernet aeussersten Operator (Wurzel) aus einem Term und schreibt 
wneuper@59523
   335
   Argumente in eine Liste *)
wneuper@59523
   336
fun args u : term list =
walther@60318
   337
  let
walther@60318
   338
    fun stripc (f $ t, ts) = stripc (f, t::ts)
walther@60318
   339
  	  | stripc (t as Free _, ts) = (t::ts)
walther@60318
   340
  	  | stripc (_, ts) = ts
walther@60318
   341
  in stripc (u, []) end;
wneuper@59523
   342
                                    
wneuper@59523
   343
(* liefert True, falls der Term (Liste von Termen) nur Zahlen 
wneuper@59523
   344
   (keine Variablen) enthaelt *)
walther@60317
   345
fun filter_num ts = fold (curry and_) (map TermC.is_num ts) true
wneuper@59523
   346
wneuper@59523
   347
(* liefert True, falls der Term nur Zahlen (keine Variablen) enthaelt 
wneuper@59523
   348
   dh. er ist ein numerischer Wert und entspricht einem Koeffizienten *)
wneuper@59523
   349
fun is_nums t = filter_num [t];
wneuper@59523
   350
wneuper@59523
   351
(* Berechnet den Gesamtgrad eines Monoms *)
walther@60318
   352
(**)local(**)
walther@60317
   353
  fun counter (n, []) = n
walther@60317
   354
    | counter (n, x :: xs) = 
walther@60317
   355
	    if (is_nums x) then counter (n, xs)
walther@60317
   356
	    else 
walther@60317
   357
	      (case x of 
wenzelm@60405
   358
		      (Const (\<^const_name>\<open>realpow\<close>, _) $ Free _ $ t) =>
walther@60318
   359
            if TermC.is_num t
walther@60318
   360
            then counter (t |> HOLogic.dest_number |> snd |> curry op + n, xs)
walther@60318
   361
            else counter (n + 1000, xs) (*FIXME.MG?!*)
walther@60335
   362
	      | (Const (\<^const_name>\<open>numeral\<close>, _) $ num) =>
walther@60318
   363
            counter (n + 1 + HOLogic.dest_numeral num, xs)
walther@60318
   364
	      | _ => counter (n + 1, xs)) (*FIXME.MG?! ... Brüche ect.*)
walther@60318
   365
(**)in(**)
walther@60317
   366
  fun monom_degree l = counter (0, l) 
walther@60318
   367
(**)end;(*local*)
wneuper@59523
   368
wneuper@59523
   369
(* wie Ordnung dict_ord (lexicographische Ordnung zweier Listen, mit Vergleich 
wneuper@59523
   370
   der Listen-Elemente mit elem_ord) - Elemente die Bedingung cond erfuellen, 
wneuper@59523
   371
   werden jedoch dabei ignoriert (uebersprungen)  *)
walther@60321
   372
fun dict_cond_ord _ _ ([], [])     = (@{print} {a = "dict_cond_ord ([], [])"}; EQUAL)
walther@60321
   373
  | dict_cond_ord _ _ ([], _ :: _) = (@{print} {a = "dict_cond_ord ([], _ :: _)"}; LESS)
walther@60321
   374
  | dict_cond_ord _ _ (_ :: _, []) = (@{print} {a = "dict_cond_ord (_ :: _, [])"}; GREATER)
walther@60321
   375
  | dict_cond_ord elem_ord cond (x :: xs, y :: ys) =
walther@60321
   376
    (@{print} {a = "dict_cond_ord", args = "(" ^ UnparseC.terms (x :: xs) ^ ", " ^ UnparseC.terms (y :: ys) ^ ")", 
walther@60321
   377
      is_nums = "(" ^ LibraryC.bool2str (cond x) ^ ", " ^ LibraryC.bool2str (cond y) ^ ")"};
walther@60321
   378
     case (cond x, cond y) of 
walther@60321
   379
	    (false, false) =>
walther@60321
   380
        (case elem_ord (x, y) of 
walther@60321
   381
				  EQUAL => dict_cond_ord elem_ord cond (xs, ys) 
walther@60321
   382
			  | ord => ord)
walther@60321
   383
    | (false, true)  => dict_cond_ord elem_ord cond (x :: xs, ys)
walther@60321
   384
    | (true, false)  => dict_cond_ord elem_ord cond (xs, y :: ys)
walther@60321
   385
    | (true, true)  =>  dict_cond_ord elem_ord cond (xs, ys) );
wneuper@59523
   386
fun dict_cond_ord _ _ ([], []) = EQUAL
wneuper@59523
   387
  | dict_cond_ord _ _ ([], _ :: _) = LESS
wneuper@59523
   388
  | dict_cond_ord _ _ (_ :: _, []) = GREATER
wneuper@59523
   389
  | dict_cond_ord elem_ord cond (x :: xs, y :: ys) =
wneuper@59523
   390
    (case (cond x, cond y) of 
walther@60318
   391
	    (false, false) =>
walther@60318
   392
        (case elem_ord (x, y) of 
walther@60318
   393
				  EQUAL => dict_cond_ord elem_ord cond (xs, ys) 
walther@60318
   394
			  | ord => ord)
walther@60318
   395
    | (false, true)  => dict_cond_ord elem_ord cond (x :: xs, ys)
walther@60318
   396
    | (true, false)  => dict_cond_ord elem_ord cond (xs, y :: ys)
walther@60318
   397
    | (true, true)  =>  dict_cond_ord elem_ord cond (xs, ys) );
wneuper@59523
   398
wneuper@59523
   399
(* Gesamtgradordnung zum Vergleich von Monomen (Liste von Variablen/Potenzen):
wneuper@59523
   400
   zuerst nach Gesamtgrad, bei gleichem Gesamtgrad lexikographisch ordnen - 
walther@60260
   401
   dabei werden Koeffizienten ignoriert (2*3*a \<up> 2*4*b gilt wie a \<up> 2*b) *)
wneuper@59523
   402
fun degree_ord (xs, ys) =
walther@60318
   403
	prod_ord int_ord (dict_cond_ord var_ord_revPow is_nums)
walther@60318
   404
	  ((monom_degree xs, xs), (monom_degree ys, ys));
wneuper@59523
   405
wneuper@59523
   406
fun hd_str str = substring (str, 0, 1);
wneuper@59523
   407
fun tl_str str = substring (str, 1, (size str) - 1);
wneuper@59523
   408
wneuper@59523
   409
(* liefert nummerischen Koeffizienten eines Monoms oder NONE *)
walther@60318
   410
fun get_koeff_of_mon [] = raise ERROR "get_koeff_of_mon: called with l = []"
walther@60318
   411
  | get_koeff_of_mon (x :: _) = if is_nums x then SOME x else NONE;
wneuper@59523
   412
wneuper@59523
   413
(* wandelt Koeffizient in (zum sortieren geeigneten) String um *)
walther@60318
   414
fun koeff2ordStr (SOME t) =
walther@60318
   415
    if TermC.is_num t
walther@60318
   416
    then 
walther@60318
   417
      if (t |> HOLogic.dest_number |> snd) < 0
walther@60318
   418
      then (t |> HOLogic.dest_number |> snd |> curry op * ~1 |> string_of_int) ^ "0"  (* 3 < -3 *)
walther@60318
   419
      else (t |> HOLogic.dest_number |> snd |> string_of_int)
walther@60318
   420
    else "aaa"                                                      (* "num.Ausdruck" --> gross *)
walther@60318
   421
  | koeff2ordStr NONE = "---";                                     (* "kein Koeff" --> kleinste *)
wneuper@59523
   422
wneuper@59523
   423
(* Order zum Vergleich von Koeffizienten (strings): 
wneuper@59523
   424
   "kein Koeff" < "0" < "1" < "-1" < "2" < "-2" < ... < "num.Ausdruck" *)
walther@60318
   425
fun compare_koeff_ord (xs, ys) = string_ord
walther@60318
   426
  ((koeff2ordStr o get_koeff_of_mon) xs,
walther@60318
   427
   (koeff2ordStr o get_koeff_of_mon) ys);
wneuper@59523
   428
wneuper@59523
   429
(* Gesamtgradordnung degree_ord + Ordnen nach Koeffizienten falls EQUAL *)
wneuper@59523
   430
fun koeff_degree_ord (xs, ys) =
wneuper@59523
   431
	    prod_ord degree_ord compare_koeff_ord ((xs, xs), (ys, ys));
wneuper@59523
   432
wneuper@59523
   433
(* Ordnet ein Liste von Monomen (Monom = Liste von Variablen) mittels 
wneuper@59523
   434
   Gesamtgradordnung *)
wneuper@59523
   435
val sort_monList = sort koeff_degree_ord;
wneuper@59523
   436
wneuper@59523
   437
(* Alternativ zu degree_ord koennte auch die viel einfachere und 
wneuper@59523
   438
   kuerzere Ordnung simple_ord verwendet werden - ist aber nicht 
wneuper@59523
   439
   fuer unsere Zwecke geeignet!
wneuper@59523
   440
wneuper@59523
   441
fun simple_ord (al,bl: term list) = dict_ord string_ord 
wneuper@59523
   442
	 (map get_basStr al, map get_basStr bl); 
wneuper@59523
   443
wneuper@59523
   444
val sort_monList = sort simple_ord; *)
wneuper@59523
   445
wneuper@59523
   446
(* aus 2 Variablen wird eine Summe bzw ein Produkt erzeugt 
wneuper@59523
   447
   (mit gewuenschtem Typen T) *)
wenzelm@60309
   448
fun plus T = Const (\<^const_name>\<open>plus\<close>, [T,T] ---> T);
wenzelm@60309
   449
fun mult T = Const (\<^const_name>\<open>times\<close>, [T,T] ---> T);
wneuper@59523
   450
fun binop op_ t1 t2 = op_ $ t1 $ t2;
wneuper@59523
   451
fun create_prod T (a,b) = binop (mult T) a b;
wneuper@59523
   452
fun create_sum T (a,b) = binop (plus T) a b;
wneuper@59523
   453
wneuper@59523
   454
(* löscht letztes Element einer Liste *)
wneuper@59523
   455
fun drop_last l = take ((length l)-1,l);
wneuper@59523
   456
wneuper@59523
   457
(* Liste von Variablen --> Monom *)
wneuper@59523
   458
fun create_monom T vl = foldr (create_prod T) (drop_last vl, last_elem vl);
wneuper@59523
   459
(* Bemerkung: 
wneuper@59523
   460
   foldr bewirkt rechtslastige Klammerung des Monoms - ist notwendig, damit zwei 
wneuper@59523
   461
   gleiche Monome zusammengefasst werden können (collect_numerals)! 
wneuper@59523
   462
   zB: 2*(x*(y*z)) + 3*(x*(y*z)) --> (2+3)*(x*(y*z))*)
wneuper@59523
   463
wneuper@59523
   464
(* Liste von Monomen --> Polynom *)	
wneuper@59523
   465
fun create_polynom T ml = foldl (create_sum T) (hd ml, tl ml);
wneuper@59523
   466
(* Bemerkung: 
wneuper@59523
   467
   foldl bewirkt linkslastige Klammerung des Polynoms (der Summanten) - 
wneuper@59523
   468
   bessere Darstellung, da keine Klammern sichtbar! 
wneuper@59523
   469
   (und discard_parentheses in make_polynomial hat weniger zu tun) *)
wneuper@59523
   470
wneuper@59523
   471
(* sorts the variables (faktors) of an expanded polynomial lexicographical *)
wneuper@59523
   472
fun sort_variables t = 
walther@60317
   473
  let
walther@60317
   474
  	val ll = map monom2list (poly2list t);
walther@60317
   475
  	val lls = map sort_varList ll; 
walther@60317
   476
  	val T = type_of t;
walther@60317
   477
  	val ls = map (create_monom T) lls;
walther@60317
   478
  in create_polynom T ls end;
wneuper@59523
   479
wneuper@59523
   480
(* sorts the monoms of an expanded and variable-sorted polynomial 
wneuper@59523
   481
   by total_degree *)
wneuper@59523
   482
fun sort_monoms t = 
walther@60318
   483
  let
walther@60318
   484
  	val ll =  map monom2list (poly2list t);
walther@60318
   485
  	val lls = sort_monList ll;
walther@60318
   486
  	val T = Term.type_of t;
walther@60318
   487
  	val ls = map (create_monom T) lls;
walther@60318
   488
  in create_polynom T ls end;
wneuper@59523
   489
\<close>
wneuper@59523
   490
wneuper@59523
   491
subsubsection \<open>rewrite order for hard-coded AC rewriting\<close>
wneuper@59523
   492
ML \<open>
neuper@37950
   493
local (*. for make_polynomial .*)
neuper@37950
   494
neuper@37950
   495
open Term;  (* for type order = EQUAL | LESS | GREATER *)
neuper@37950
   496
neuper@37950
   497
fun pr_ord EQUAL = "EQUAL"
neuper@37950
   498
  | pr_ord LESS  = "LESS"
neuper@37950
   499
  | pr_ord GREATER = "GREATER";
neuper@37950
   500
neuper@37950
   501
fun dest_hd' (Const (a, T)) =                          (* ~ term.ML *)
neuper@37950
   502
  (case a of
wenzelm@60405
   503
     \<^const_name>\<open>realpow\<close> => ((("|||||||||||||", 0), T), 0)    (*WN greatest string*)
neuper@37950
   504
   | _ => (((a, 0), T), 0))
walther@60317
   505
  | dest_hd' (Free (a, T)) = (((a, 0), T), 1)(*TODOO handle this as numeral, too? see EqSystem.thy*)
neuper@37950
   506
  | dest_hd' (Var v) = (v, 2)
neuper@37950
   507
  | dest_hd' (Bound i) = ((("", i), dummyT), 3)
wneuper@59523
   508
  | dest_hd' (Abs (_, T, _)) = ((("", 0), T), 4)
wneuper@59523
   509
  | dest_hd' t = raise TERM ("dest_hd'", [t]);
neuper@37950
   510
neuper@37950
   511
fun size_of_term' (Const(str,_) $ t) =
wenzelm@60405
   512
  if \<^const_name>\<open>realpow\<close>= str then 1000 + size_of_term' t else 1+size_of_term' t(*WN*)
neuper@37950
   513
  | size_of_term' (Abs (_,_,body)) = 1 + size_of_term' body
neuper@37950
   514
  | size_of_term' (f$t) = size_of_term' f  +  size_of_term' t
neuper@37950
   515
  | size_of_term' _ = 1;
neuper@37950
   516
neuper@37950
   517
fun term_ord' pr thy (Abs (_, T, t), Abs(_, U, u)) =       (* ~ term.ML *)
neuper@52070
   518
    (case term_ord' pr thy (t, u) of EQUAL => Term_Ord.typ_ord (T, U) | ord => ord)
neuper@37950
   519
  | term_ord' pr thy (t, u) =
neuper@52070
   520
    (if pr then 
neuper@52070
   521
	   let
neuper@52070
   522
       val (f, ts) = strip_comb t and (g, us) = strip_comb u;
walther@59870
   523
       val _ = tracing ("t= f@ts= \"" ^ UnparseC.term_in_thy thy f ^ "\" @ \"[" ^
walther@59870
   524
         commas (map (UnparseC.term_in_thy thy) ts) ^ "]\"");
walther@59870
   525
       val _ = tracing("u= g@us= \"" ^ UnparseC.term_in_thy thy g ^ "\" @ \"[" ^
walther@59870
   526
         commas (map (UnparseC.term_in_thy thy) us) ^ "]\"");
neuper@52070
   527
       val _ = tracing ("size_of_term(t,u)= (" ^ string_of_int (size_of_term' t) ^ ", " ^
neuper@52070
   528
         string_of_int (size_of_term' u) ^ ")");
neuper@52070
   529
       val _ = tracing ("hd_ord(f,g)      = " ^ (pr_ord o hd_ord) (f,g));
neuper@52070
   530
       val _ = tracing ("terms_ord(ts,us) = " ^ (pr_ord o terms_ord str false) (ts, us));
neuper@52070
   531
       val _ = tracing ("-------");
neuper@52070
   532
     in () end
neuper@37950
   533
       else ();
neuper@37950
   534
	 case int_ord (size_of_term' t, size_of_term' u) of
neuper@37950
   535
	   EQUAL =>
neuper@37950
   536
	     let val (f, ts) = strip_comb t and (g, us) = strip_comb u in
neuper@37950
   537
	       (case hd_ord (f, g) of EQUAL => (terms_ord str pr) (ts, us) 
neuper@37950
   538
	     | ord => ord)
neuper@37950
   539
	     end
neuper@37950
   540
	 | ord => ord)
neuper@37950
   541
and hd_ord (f, g) =                                        (* ~ term.ML *)
neuper@37974
   542
  prod_ord (prod_ord Term_Ord.indexname_ord Term_Ord.typ_ord) int_ord (dest_hd' f, dest_hd' g)
wneuper@59523
   543
and terms_ord _ pr (ts, us) = 
walther@59881
   544
    list_ord (term_ord' pr (ThyC.get_theory "Isac_Knowledge"))(ts, us);
neuper@52070
   545
neuper@37950
   546
in
neuper@37950
   547
walther@60324
   548
fun ord_make_polynomial (pr:bool) thy (_: subst) (ts, us) =
walther@60324
   549
    (term_ord' pr thy(***) (TermC.numerals_to_Free ts, TermC.numerals_to_Free us) = LESS );
neuper@37950
   550
neuper@37950
   551
end;(*local*)
Walther@60506
   552
\<close> 
Walther@60506
   553
setup \<open>KEStore_Elems.add_rew_ord [
Walther@60506
   554
  ("termlessI", termlessI), 
Walther@60506
   555
  ("ord_make_polynomial", ord_make_polynomial false \<^theory>)]\<close>
neuper@37950
   556
wneuper@59523
   557
subsection \<open>predicates\<close>
wneuper@59523
   558
subsubsection \<open>in specifications\<close>
wneuper@59523
   559
ML \<open>
wneuper@59523
   560
(* is_polyrat_in becomes true, if no bdv is in the denominator of a fraction*)
wneuper@59523
   561
fun is_polyrat_in t v = 
wneuper@59524
   562
  let
walther@59962
   563
   	fun finddivide (_ $ _ $ _ $ _) _ = raise ERROR("is_polyrat_in:")
wneuper@59523
   564
	    (* at the moment there is no term like this, but ....*)
wenzelm@60309
   565
	  | finddivide (Const (\<^const_name>\<open>divide\<close>,_) $ _ $ b) v = not (Prog_Expr.occurs_in v b)
wneuper@59524
   566
	  | finddivide (_ $ t1 $ t2) v = finddivide t1 v orelse finddivide t2 v
wneuper@59524
   567
	  | finddivide (_ $ t1) v = finddivide t1 v
wneuper@59523
   568
	  | finddivide _ _ = false;
wneuper@59524
   569
  in finddivide t v end;
wneuper@59523
   570
    
wneuper@59524
   571
fun is_expanded_in t v = case expand_deg_in t v of SOME _ => true | NONE => false;
wneuper@59524
   572
fun is_poly_in t v =     case poly_deg_in t v of SOME _ => true | NONE => false;
wneuper@59524
   573
fun has_degree_in t v =  case expand_deg_in t v of SOME d => d | NONE => ~1;
neuper@37950
   574
wneuper@59523
   575
(*.the expression contains + - * ^ only ?
wneuper@59523
   576
   this is weaker than 'is_polynomial' !.*)
wneuper@59523
   577
fun is_polyexp (Free _) = true
wneuper@59529
   578
  | is_polyexp (Const _) = true (* potential danger: bdv is not considered *)
walther@60335
   579
  | is_polyexp (Const (\<^const_name>\<open>plus\<close>,_) $ Free _ $ num) =
walther@60318
   580
    if TermC.is_num num then true
walther@60318
   581
    else if TermC.is_variable num then true
walther@60318
   582
    else is_polyexp num
walther@60335
   583
  | is_polyexp (Const (\<^const_name>\<open>plus\<close>, _) $ num $ Free _) =
walther@60318
   584
    if TermC.is_num num then true
walther@60318
   585
    else if TermC.is_variable num then true
walther@60318
   586
    else is_polyexp num
walther@60335
   587
  | is_polyexp (Const (\<^const_name>\<open>minus\<close>, _) $ Free _ $ num) =
walther@60318
   588
    if TermC.is_num num then true
walther@60318
   589
    else if TermC.is_variable num then true
walther@60318
   590
    else is_polyexp num
walther@60335
   591
  | is_polyexp (Const (\<^const_name>\<open>times\<close>, _) $ num $ Free _) =
walther@60318
   592
    if TermC.is_num num then true
walther@60318
   593
    else if TermC.is_variable num then true
walther@60318
   594
    else is_polyexp num
wenzelm@60405
   595
  | is_polyexp (Const (\<^const_name>\<open>realpow\<close>,_) $ Free _ $ num) =
walther@60318
   596
    if TermC.is_num num then true
walther@60318
   597
    else if TermC.is_variable num then true
walther@60318
   598
    else is_polyexp num
walther@60335
   599
  | is_polyexp (Const (\<^const_name>\<open>plus_class.plus\<close>,_) $ t1 $ t2) = 
walther@60318
   600
    ((is_polyexp t1) andalso (is_polyexp t2))
walther@60335
   601
  | is_polyexp (Const (\<^const_name>\<open>Groups.minus_class.minus\<close>,_) $ t1 $ t2) = 
walther@60318
   602
    ((is_polyexp t1) andalso (is_polyexp t2))
walther@60335
   603
  | is_polyexp (Const (\<^const_name>\<open>Groups.times_class.times\<close>,_) $ t1 $ t2) = 
walther@60318
   604
    ((is_polyexp t1) andalso (is_polyexp t2))
wenzelm@60405
   605
  | is_polyexp (Const (\<^const_name>\<open>realpow\<close>,_) $ t1 $ t2) = 
walther@60318
   606
    ((is_polyexp t1) andalso (is_polyexp t2))
walther@60317
   607
  | is_polyexp num = TermC.is_num num;
wneuper@59523
   608
\<close>
neuper@37950
   609
wneuper@59523
   610
subsubsection \<open>for hard-coded AC rewriting\<close>
wneuper@59523
   611
ML \<open>
wneuper@59523
   612
(* auch Klammerung muss übereinstimmen;
wneuper@59523
   613
   sort_variables klammert Produkte rechtslastig*)
wneuper@59523
   614
fun is_multUnordered t = ((is_polyexp t) andalso not (t = sort_variables t));
wneuper@59523
   615
wneuper@59523
   616
fun is_addUnordered t = ((is_polyexp t) andalso not (t = sort_monoms t));
wneuper@59523
   617
\<close>
wneuper@59523
   618
wneuper@59523
   619
subsection \<open>evaluations functions\<close>
walther@60278
   620
subsubsection \<open>for predicates\<close>
wneuper@59523
   621
ML \<open>
walther@60335
   622
fun eval_is_polyrat_in _ _(p as (Const (\<^const_name>\<open>Poly.is_polyrat_in\<close>, _) $ t $ v)) _  =
wneuper@59523
   623
    if is_polyrat_in t v 
walther@59868
   624
    then SOME ((UnparseC.term p) ^ " = True",
wneuper@59523
   625
	        HOLogic.Trueprop $ (TermC.mk_equality (p, @{term True})))
walther@59868
   626
    else SOME ((UnparseC.term p) ^ " = True",
wneuper@59523
   627
	        HOLogic.Trueprop $ (TermC.mk_equality (p, @{term False})))
wneuper@59523
   628
  | eval_is_polyrat_in _ _ _ _ = ((*tracing"### no matches";*) NONE);
wneuper@59523
   629
walther@60278
   630
(*("is_expanded_in", ("Poly.is_expanded_in", eval_is_expanded_in ""))*)
wneuper@59523
   631
fun eval_is_expanded_in _ _ 
walther@60335
   632
       (p as (Const (\<^const_name>\<open>Poly.is_expanded_in\<close>, _) $ t $ v)) _ =
wneuper@59523
   633
    if is_expanded_in t v
walther@59868
   634
    then SOME ((UnparseC.term p) ^ " = True",
wneuper@59523
   635
	        HOLogic.Trueprop $ (TermC.mk_equality (p, @{term True})))
walther@59868
   636
    else SOME ((UnparseC.term p) ^ " = True",
wneuper@59523
   637
	        HOLogic.Trueprop $ (TermC.mk_equality (p, @{term False})))
wneuper@59523
   638
  | eval_is_expanded_in _ _ _ _ = NONE;
wneuper@59523
   639
walther@60278
   640
(*("is_poly_in", ("Poly.is_poly_in", eval_is_poly_in ""))*)
wneuper@59523
   641
fun eval_is_poly_in _ _ 
walther@60335
   642
       (p as (Const (\<^const_name>\<open>Poly.is_poly_in\<close>, _) $ t $ v)) _ =
wneuper@59523
   643
    if is_poly_in t v
walther@59868
   644
    then SOME ((UnparseC.term p) ^ " = True",
wneuper@59523
   645
	        HOLogic.Trueprop $ (TermC.mk_equality (p, @{term True})))
walther@59868
   646
    else SOME ((UnparseC.term p) ^ " = True",
wneuper@59523
   647
	        HOLogic.Trueprop $ (TermC.mk_equality (p, @{term False})))
wneuper@59523
   648
  | eval_is_poly_in _ _ _ _ = NONE;
wneuper@59523
   649
walther@60278
   650
(*("has_degree_in", ("Poly.has_degree_in", eval_has_degree_in ""))*)
wneuper@59523
   651
fun eval_has_degree_in _ _ 
walther@60335
   652
	     (p as (Const (\<^const_name>\<open>Poly.has_degree_in\<close>, _) $ t $ v)) _ =
wneuper@59523
   653
    let val d = has_degree_in t v
wneuper@59523
   654
	val d' = TermC.term_of_num HOLogic.realT d
walther@59868
   655
    in SOME ((UnparseC.term p) ^ " = " ^ (string_of_int d),
wneuper@59523
   656
	      HOLogic.Trueprop $ (TermC.mk_equality (p, d')))
wneuper@59523
   657
    end
wneuper@59523
   658
  | eval_has_degree_in _ _ _ _ = NONE;
wneuper@59523
   659
walther@60278
   660
(*("is_polyexp", ("Poly.is_polyexp", eval_is_polyexp ""))*)
wneuper@59523
   661
fun eval_is_polyexp (thmid:string) _ 
Walther@60504
   662
		       (t as (Const (\<^const_name>\<open>is_polyexp\<close>, _) $ arg)) ctxt = 
wneuper@59523
   663
    if is_polyexp arg
Walther@60504
   664
    then SOME (TermC.mk_thmid thmid (UnparseC.term_in_ctxt ctxt arg) "", 
wneuper@59523
   665
	         HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True})))
Walther@60504
   666
    else SOME (TermC.mk_thmid thmid (UnparseC.term_in_ctxt ctxt arg) "", 
wneuper@59523
   667
	         HOLogic.Trueprop $ (TermC.mk_equality (t, @{term False})))
wneuper@59523
   668
  | eval_is_polyexp _ _ _ _ = NONE; 
wneuper@59523
   669
\<close>
wneuper@59523
   670
wneuper@59523
   671
subsubsection \<open>for hard-coded AC rewriting\<close>
wneuper@59523
   672
ML \<open>
wneuper@59523
   673
(*WN.18.6.03 *)
walther@60278
   674
(*("is_addUnordered", ("Poly.is_addUnordered", eval_is_addUnordered ""))*)
wneuper@59523
   675
fun eval_is_addUnordered (thmid:string) _ 
Walther@60504
   676
		       (t as (Const (\<^const_name>\<open>is_addUnordered\<close>, _) $ arg)) ctxt = 
wneuper@59523
   677
    if is_addUnordered arg
Walther@60504
   678
    then SOME (TermC.mk_thmid thmid (UnparseC.term_in_ctxt ctxt arg) "", 
wneuper@59523
   679
	         HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True})))
Walther@60504
   680
    else SOME (TermC.mk_thmid thmid (UnparseC.term_in_ctxt ctxt arg) "", 
wneuper@59523
   681
	         HOLogic.Trueprop $ (TermC.mk_equality (t, @{term False})))
wneuper@59523
   682
  | eval_is_addUnordered _ _ _ _ = NONE; 
wneuper@59523
   683
wneuper@59523
   684
fun eval_is_multUnordered (thmid:string) _ 
Walther@60504
   685
		       (t as (Const (\<^const_name>\<open>is_multUnordered\<close>, _) $ arg)) ctxt = 
wneuper@59523
   686
    if is_multUnordered arg
Walther@60504
   687
    then SOME (TermC.mk_thmid thmid (UnparseC.term_in_ctxt ctxt arg) "", 
wneuper@59523
   688
	         HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True})))
Walther@60504
   689
    else SOME (TermC.mk_thmid thmid (UnparseC.term_in_ctxt ctxt arg) "", 
wneuper@59523
   690
	         HOLogic.Trueprop $ (TermC.mk_equality (t, @{term False})))
wneuper@59523
   691
  | eval_is_multUnordered _ _ _ _ = NONE; 
wneuper@59523
   692
\<close>
wenzelm@60313
   693
calculation is_polyrat_in = \<open>eval_is_polyrat_in "#eval_is_polyrat_in"\<close>
wenzelm@60313
   694
calculation is_expanded_in = \<open>eval_is_expanded_in ""\<close>
wenzelm@60313
   695
calculation is_poly_in = \<open>eval_is_poly_in ""\<close>
wenzelm@60313
   696
calculation has_degree_in = \<open>eval_has_degree_in ""\<close>
wenzelm@60313
   697
calculation is_polyexp = \<open>eval_is_polyexp ""\<close>
wenzelm@60313
   698
calculation is_multUnordered = \<open>eval_is_multUnordered ""\<close>
wenzelm@60313
   699
calculation is_addUnordered = \<open>eval_is_addUnordered ""\<close>
wneuper@59523
   700
wneuper@59523
   701
subsection \<open>rule-sets\<close>
wneuper@59523
   702
subsubsection \<open>without specific order\<close>
wneuper@59523
   703
ML \<open>
wneuper@59523
   704
(* used only for merge *)
walther@59852
   705
val calculate_Poly = Rule_Set.append_rules "calculate_PolyFIXXXME.not.impl." Rule_Set.empty [];
wneuper@59523
   706
wneuper@59523
   707
(*.for evaluation of conditions in rewrite rules.*)
walther@60358
   708
val Poly_erls = Rule_Set.append_rules "Poly_erls" Atools_erls [
walther@60358
   709
  \<^rule_eval>\<open>HOL.eq\<close> (Prog_Expr.eval_equal "#equal_"),
wenzelm@60297
   710
  \<^rule_thm>\<open>real_unari_minus\<close>,
wenzelm@60294
   711
  \<^rule_eval>\<open>plus\<close> (eval_binop "#add_"),
wenzelm@60294
   712
  \<^rule_eval>\<open>minus\<close> (eval_binop "#sub_"),
wenzelm@60294
   713
  \<^rule_eval>\<open>times\<close> (eval_binop "#mult_"),
wenzelm@60405
   714
  \<^rule_eval>\<open>realpow\<close> (eval_binop "#power_")];
wneuper@59523
   715
walther@60358
   716
val poly_crls = Rule_Set.append_rules "poly_crls" Atools_crls [
walther@60358
   717
  \<^rule_eval>\<open>HOL.eq\<close> (Prog_Expr.eval_equal "#equal_"),
wenzelm@60297
   718
  \<^rule_thm>\<open>real_unari_minus\<close>,
wenzelm@60294
   719
  \<^rule_eval>\<open>plus\<close> (eval_binop "#add_"),
wenzelm@60294
   720
  \<^rule_eval>\<open>minus\<close> (eval_binop "#sub_"),
wenzelm@60294
   721
  \<^rule_eval>\<open>times\<close> (eval_binop "#mult_"),
wenzelm@60405
   722
  \<^rule_eval>\<open>realpow\<close> (eval_binop "#power_")];
wneuper@59523
   723
\<close>
wneuper@59523
   724
ML \<open>
neuper@37950
   725
val expand =
Walther@60509
   726
  Rule_Def.Repeat {id = "expand", preconds = [], rew_ord = ("dummy_ord", Rewrite_Ord.function_empty),
walther@59852
   727
      erls = Rule_Set.empty,srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@60358
   728
      rules = [
walther@60358
   729
        \<^rule_thm>\<open>distrib_right\<close>, (*"(z1.0 + z2.0) * w = z1.0 * w + z2.0 * w"*)
walther@60358
   730
	      \<^rule_thm>\<open>distrib_left\<close> (*"w * (z1.0 + z2.0) = w * z1.0 + w * z2.0"*)],
walther@60358
   731
      scr = Rule.Empty_Prog};
neuper@37950
   732
walther@60385
   733
\<close> ML \<open>
walther@60385
   734
\<close> ML \<open>@{term "1 is_num"}\<close>
walther@60385
   735
ML \<open>
walther@60320
   736
(* erls for calculate_Rational + etc *)
walther@60318
   737
val powers_erls =
Walther@60509
   738
  Rule_Def.Repeat {id = "powers_erls", preconds = [], rew_ord = ("dummy_ord",Rewrite_Ord.function_empty), 
walther@60318
   739
      erls = Rule_Set.empty, srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@60358
   740
      rules = [
walther@60358
   741
        \<^rule_eval>\<open>matches\<close> (Prog_Expr.eval_matches "#matches_"),
walther@60358
   742
        \<^rule_eval>\<open>is_atom\<close> (Prog_Expr.eval_is_atom "#is_atom_"),
walther@60358
   743
        \<^rule_eval>\<open>is_num\<close> (Prog_Expr.eval_is_num "#is_num_"),
walther@60358
   744
        \<^rule_eval>\<open>is_even\<close> (Prog_Expr.eval_is_even "#is_even_"),
walther@60358
   745
        \<^rule_eval>\<open>ord_class.less\<close> (Prog_Expr.eval_equ "#less_"),
walther@60358
   746
        \<^rule_thm>\<open>not_false\<close>,
walther@60358
   747
        \<^rule_thm>\<open>not_true\<close>,
walther@60358
   748
        \<^rule_eval>\<open>plus_class.plus\<close> (eval_binop "#add_")],
walther@60358
   749
      scr = Rule.Empty_Prog};
walther@60320
   750
walther@60333
   751
\<close> ML \<open>
walther@60320
   752
val discard_minus =
Walther@60509
   753
  Rule_Def.Repeat {id = "discard_minus", preconds = [], rew_ord = ("dummy_ord", Rewrite_Ord.function_empty),
walther@60320
   754
      erls = powers_erls, srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@60341
   755
      rules = [
walther@60341
   756
        \<^rule_thm>\<open>real_diff_minus\<close> (*"a - b = a + -1 * b"*),
walther@60385
   757
        \<^rule_thm>\<open>real_mult_minus1_sym\<close> (*"\<not>(z is_num) ==> - (z::real) = -1 * z"*)],
walther@60341
   758
	    scr = Rule.Empty_Prog};
walther@60320
   759
neuper@37950
   760
val expand_poly_ = 
walther@59851
   761
  Rule_Def.Repeat{id = "expand_poly_", preconds = [], 
Walther@60509
   762
      rew_ord = ("dummy_ord", Rewrite_Ord.function_empty),
walther@60318
   763
      erls = powers_erls, srls = Rule_Set.Empty,
neuper@42451
   764
      calc = [], errpatts = [],
walther@60358
   765
      rules = [
walther@60358
   766
        \<^rule_thm>\<open>real_plus_binom_pow4\<close>, (*"(a + b) \<up> 4 = ... "*)
walther@60358
   767
        \<^rule_thm>\<open>real_plus_binom_pow5\<close>, (*"(a + b) \<up> 5 = ... "*)
walther@60358
   768
        \<^rule_thm>\<open>real_plus_binom_pow3\<close>, (*"(a + b) \<up> 3 = a \<up> 3 + 3*a \<up> 2*b + 3*a*b \<up> 2 + b \<up> 3" *)
walther@60358
   769
      (*WN071229 changed/removed for Schaerding -----vvv*)
walther@60358
   770
        (*\<^rule_thm>\<open>real_plus_binom_pow2\<close>,*) (*"(a + b) \<up> 2 = a \<up> 2 + 2*a*b + b \<up> 2"*)
walther@60358
   771
        \<^rule_thm>\<open>real_plus_binom_pow2\<close>, (*"(a + b) \<up> 2 = (a + b) * (a + b)"*)
walther@60358
   772
        (*\<^rule_thm>\<open>real_plus_minus_binom1_p_p\<close>,*) (*"(a + b)*(a + -1 * b) = a \<up> 2 + -1*b \<up> 2"*)
walther@60358
   773
        (*\<^rule_thm>\<open>real_plus_minus_binom2_p_p\<close>,*) (*"(a + -1 * b)*(a + b) = a \<up> 2 + -1*b \<up> 2"*)
walther@60358
   774
      (*WN071229 changed/removed for Schaerding -----^^^*)
walther@60358
   775
walther@60358
   776
        \<^rule_thm>\<open>distrib_right\<close>, (*"(z1.0 + z2.0) * w = z1.0 * w + z2.0 * w"*)
walther@60358
   777
        \<^rule_thm>\<open>distrib_left\<close>, (*"w * (z1.0 + z2.0) = w * z1.0 + w * z2.0"*)
walther@60358
   778
walther@60358
   779
        \<^rule_thm>\<open>realpow_multI\<close>, (*"(r * s) \<up> n = r \<up> n * s \<up> n"*)
walther@60358
   780
        \<^rule_thm>\<open>realpow_pow\<close>, (*"(a \<up> b) \<up> c = a \<up> (b * c)"*)
walther@60358
   781
walther@60358
   782
        \<^rule_thm>\<open>realpow_minus_even\<close>, (*"n is_even ==> (- r) \<up> n = r \<up> n"*)
walther@60358
   783
        \<^rule_thm>\<open>realpow_minus_odd\<close> (*"Not (n is_even) ==> (- r) \<up> n = -1 * r \<up> n"*)],
walther@60358
   784
      scr = Rule.Empty_Prog};
neuper@37950
   785
neuper@37950
   786
val expand_poly_rat_ = 
walther@59851
   787
  Rule_Def.Repeat{id = "expand_poly_rat_", preconds = [], 
Walther@60509
   788
      rew_ord = ("dummy_ord", Rewrite_Ord.function_empty),
walther@60358
   789
      erls =  Rule_Set.append_rules "Rule_Set.empty-expand_poly_rat_" Rule_Set.empty [
walther@60358
   790
        \<^rule_eval>\<open>is_polyexp\<close> (eval_is_polyexp ""),
walther@60358
   791
        \<^rule_eval>\<open>is_even\<close> (Prog_Expr.eval_is_even "#is_even_"),
walther@60358
   792
        \<^rule_thm>\<open>not_false\<close>,
walther@60358
   793
        \<^rule_thm>\<open>not_true\<close> ],
walther@59851
   794
      srls = Rule_Set.Empty,
neuper@42451
   795
      calc = [], errpatts = [],
walther@60358
   796
      rules = [
walther@60358
   797
        \<^rule_thm>\<open>real_plus_binom_pow4_poly\<close>, (*"[| a is_polyexp; b is_polyexp |] ==> (a + b) \<up> 4 = ... "*)
walther@60358
   798
        \<^rule_thm>\<open>real_plus_binom_pow5_poly\<close>, (*"[| a is_polyexp; b is_polyexp |] ==> (a + b) \<up> 5 = ... "*)
walther@60358
   799
        \<^rule_thm>\<open>real_plus_binom_pow2_poly\<close>, (*"[| a is_polyexp; b is_polyexp |] ==> (a + b) \<up> 2 = a \<up> 2 + 2*a*b + b \<up> 2"*)
walther@60358
   800
      	 \<^rule_thm>\<open>real_plus_binom_pow3_poly\<close>,
walther@60358
   801
      	   (*"[| a is_polyexp; b is_polyexp |] ==> (a + b) \<up> 3 = a \<up> 3 + 3*a \<up> 2*b + 3*a*b \<up> 2 + b \<up> 3" *)
walther@60358
   802
      	 \<^rule_thm>\<open>real_plus_minus_binom1_p_p\<close>, (*"(a + b)*(a + -1 * b) = a \<up> 2 + -1*b \<up> 2"*)
walther@60358
   803
      	 \<^rule_thm>\<open>real_plus_minus_binom2_p_p\<close>, (*"(a + -1 * b)*(a + b) = a \<up> 2 + -1*b \<up> 2"*)
walther@60358
   804
      	      
walther@60358
   805
      	 \<^rule_thm>\<open>real_add_mult_distrib_poly\<close>, (*"w is_polyexp ==> (z1.0 + z2.0) * w = z1.0 * w + z2.0 * w"*)
walther@60358
   806
      	 \<^rule_thm>\<open>real_add_mult_distrib2_poly\<close>, (*"w is_polyexp ==> w * (z1.0 + z2.0) = w * z1.0 + w * z2.0"*)
walther@60358
   807
      	       
walther@60358
   808
      	 \<^rule_thm>\<open>realpow_multI_poly\<close>, (*"[| r is_polyexp; s is_polyexp |] ==> (r * s) \<up> n = r \<up> n * s \<up> n"*)
walther@60358
   809
      	 \<^rule_thm>\<open>realpow_pow\<close>, (*"(a \<up> b) \<up> c = a \<up> (b * c)"*)
walther@60358
   810
      	 \<^rule_thm>\<open>realpow_minus_even\<close>, (*"n is_even ==> (- r) \<up> n = r \<up> n"*)
walther@60358
   811
     	   \<^rule_thm>\<open>realpow_minus_odd\<close> (*"\<not> (n is_even) ==> (- r) \<up> n = -1 * r \<up> n"*) ],
walther@60358
   812
      scr = Rule.Empty_Prog};
neuper@37950
   813
neuper@37950
   814
val simplify_power_ = 
walther@59851
   815
  Rule_Def.Repeat{id = "simplify_power_", preconds = [], 
Walther@60509
   816
      rew_ord = ("dummy_ord", Rewrite_Ord.function_empty),
walther@59852
   817
      erls = Rule_Set.empty, srls = Rule_Set.Empty,
neuper@42451
   818
      calc = [], errpatts = [],
walther@60358
   819
      rules = [
walther@60358
   820
         (*MG: Reihenfolge der folgenden 2 Rule.Thm muss so bleiben, wegen
walther@60358
   821
	        	a*(a*a) --> a*a \<up> 2 und nicht a*(a*a) --> a \<up> 2*a *)
walther@60358
   822
        \<^rule_thm_sym>\<open>realpow_twoI\<close>, (*"r * r = r \<up> 2"*)
walther@60358
   823
        \<^rule_thm>\<open>realpow_twoI_assoc_l\<close>, (*"r * (r * s) = r \<up> 2 * s"*)
walther@60358
   824
        
walther@60358
   825
        \<^rule_thm>\<open>realpow_plus_1\<close>, (*"r * r \<up> n = r \<up> (n + 1)"*)
walther@60358
   826
        \<^rule_thm>\<open>realpow_plus_1_assoc_l\<close>, (*"r * (r \<up> m * s) = r \<up> (1 + m) * s"*)
walther@60358
   827
        (*MG 9.7.03: neues Rule.Thm wegen a*(a*(a*b)) --> a \<up> 2*(a*b) *)
walther@60358
   828
        \<^rule_thm>\<open>realpow_plus_1_assoc_l2\<close>, (*"r \<up> m * (r * s) = r \<up> (1 + m) * s"*)
walther@60358
   829
        
walther@60358
   830
        \<^rule_thm_sym>\<open>realpow_addI\<close>, (*"r \<up> n * r \<up> m = r \<up> (n + m)"*)
walther@60358
   831
        \<^rule_thm>\<open>realpow_addI_assoc_l\<close>, (*"r \<up> n * (r \<up> m * s) = r \<up> (n + m) * s"*)
walther@60358
   832
        
walther@60358
   833
        (* ist in expand_poly - wird hier aber auch gebraucht, wegen: "r * r = r \<up> 2" wenn r=a \<up> b*)
walther@60358
   834
        \<^rule_thm>\<open>realpow_pow\<close> (*"(a \<up> b) \<up> c = a \<up> (b * c)"*)], 
walther@60358
   835
      scr = Rule.Empty_Prog};
neuper@37950
   836
neuper@37950
   837
val calc_add_mult_pow_ = 
walther@59851
   838
  Rule_Def.Repeat{id = "calc_add_mult_pow_", preconds = [], 
Walther@60509
   839
      rew_ord = ("dummy_ord", Rewrite_Ord.function_empty),
walther@59851
   840
      erls = Atools_erls(*erls3.4.03*),srls = Rule_Set.Empty,
wenzelm@60309
   841
      calc = [("PLUS"  , (\<^const_name>\<open>plus\<close>, eval_binop "#add_")), 
wenzelm@60309
   842
	      ("TIMES" , (\<^const_name>\<open>times\<close>, eval_binop "#mult_")),
wenzelm@60405
   843
	      ("POWER", (\<^const_name>\<open>realpow\<close>, eval_binop "#power_"))
neuper@37950
   844
	      ],
neuper@42451
   845
      errpatts = [],
walther@60358
   846
      rules = [
walther@60358
   847
        \<^rule_eval>\<open>plus\<close> (eval_binop "#add_"),
walther@60358
   848
	      \<^rule_eval>\<open>times\<close> (eval_binop "#mult_"),
wenzelm@60405
   849
	      \<^rule_eval>\<open>realpow\<close> (eval_binop "#power_")],
walther@60358
   850
      scr = Rule.Empty_Prog};
neuper@37950
   851
neuper@37950
   852
val reduce_012_mult_ = 
walther@59851
   853
  Rule_Def.Repeat{id = "reduce_012_mult_", preconds = [], 
Walther@60509
   854
      rew_ord = ("dummy_ord", Rewrite_Ord.function_empty),
walther@59852
   855
      erls = Rule_Set.empty,srls = Rule_Set.Empty,
neuper@42451
   856
      calc = [], errpatts = [],
wneuper@59416
   857
      rules = [(* MG: folgende Rule.Thm müssen hier stehen bleiben: *)
walther@60358
   858
         \<^rule_thm>\<open>mult_1_right\<close>, (*"z * 1 = z"*) (*wegen "a * b * b \<up> (-1) + a"*) 
walther@60358
   859
	       \<^rule_thm>\<open>realpow_zeroI\<close>, (*"r \<up> 0 = 1"*) (*wegen "a*a \<up> (-1)*c + b + c"*)
walther@60358
   860
	       \<^rule_thm>\<open>realpow_oneI\<close>, (*"r \<up> 1 = r"*)
walther@60358
   861
	       \<^rule_thm>\<open>realpow_eq_oneI\<close> (*"1 \<up> n = 1"*)],
walther@60358
   862
      scr = Rule.Empty_Prog};
neuper@37950
   863
neuper@37950
   864
val collect_numerals_ = 
walther@59851
   865
  Rule_Def.Repeat{id = "collect_numerals_", preconds = [], 
Walther@60509
   866
      rew_ord = ("dummy_ord", Rewrite_Ord.function_empty),
walther@59851
   867
      erls = Atools_erls, srls = Rule_Set.Empty,
wenzelm@60309
   868
      calc = [("PLUS"  , (\<^const_name>\<open>plus\<close>, eval_binop "#add_"))
neuper@42451
   869
	      ], errpatts = [],
walther@60358
   870
      rules = [
walther@60385
   871
        \<^rule_thm>\<open>real_num_collect\<close>, (*"[| l is_num; m is_num |]==>l * n + m * n = (l + m) * n"*)
walther@60358
   872
	      \<^rule_thm>\<open>real_num_collect_assoc_r\<close>,
walther@60385
   873
	        (*"[| l is_num; m is_num |] ==> (k + m * n) + l * n = k + (l + m)*n"*)
walther@60385
   874
        \<^rule_thm>\<open>real_one_collect\<close>, (*"m is_num ==> n + m * n = (1 + m) * n"*)
walther@60385
   875
        \<^rule_thm>\<open>real_one_collect_assoc_r\<close>, (*"m is_num ==> (k + n) + m * n = k + (m + 1) * n"*)
neuper@37950
   876
walther@60358
   877
        \<^rule_eval>\<open>plus\<close> (eval_binop "#add_"),
neuper@37950
   878
walther@60358
   879
      	 (*MG: Reihenfolge der folgenden 2 Rule.Thm muss so bleiben, wegen
walther@60358
   880
		       (a+a)+a --> a + 2*a --> 3*a and not (a+a)+a --> 2*a + a *)
walther@60358
   881
        \<^rule_thm>\<open>real_mult_2_assoc_r\<close>, (*"(k + z1) + z1 = k + 2 * z1"*)
walther@60358
   882
      	 \<^rule_thm_sym>\<open>real_mult_2\<close> (*"z1 + z1 = 2 * z1"*)],
walther@60358
   883
      scr = Rule.Empty_Prog};
neuper@37950
   884
neuper@37950
   885
val reduce_012_ = 
walther@59851
   886
  Rule_Def.Repeat{id = "reduce_012_", preconds = [], 
Walther@60509
   887
      rew_ord = ("dummy_ord", Rewrite_Ord.function_empty),
walther@59852
   888
      erls = Rule_Set.empty,srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@60358
   889
      rules = [
walther@60358
   890
        \<^rule_thm>\<open>mult_1_left\<close>, (*"1 * z = z"*)
walther@60358
   891
	      \<^rule_thm>\<open>mult_zero_left\<close>, (*"0 * z = 0"*)
walther@60358
   892
	      \<^rule_thm>\<open>mult_zero_right\<close>, (*"z * 0 = 0"*)
walther@60358
   893
	      \<^rule_thm>\<open>add_0_left\<close>, (*"0 + z = z"*)
walther@60358
   894
	      \<^rule_thm>\<open>add_0_right\<close>, (*"z + 0 = z"*) (*wegen a+b-b --> a+(1-1)*b --> a+0 --> a*)
walther@60358
   895
   
walther@60358
   896
	      (*\<^rule_thm>\<open>realpow_oneI\<close>*) (*"?r \<up> 1 = ?r"*)
walther@60358
   897
	      \<^rule_thm>\<open>division_ring_divide_zero\<close> (*"0 / ?x = 0"*)],
walther@60358
   898
      scr = Rule.Empty_Prog};
neuper@37950
   899
neuper@37979
   900
val discard_parentheses1 = 
walther@60358
   901
    Rule_Set.append_rules "discard_parentheses1" Rule_Set.empty [
walther@60358
   902
      \<^rule_thm_sym>\<open>mult.assoc\<close> (*"?z1.1 * (?z2.1 * ?z3.1) = ?z1.1 * ?z2.1 * ?z3.1"*)
walther@60358
   903
		   (*\<^rule_thm_sym>\<open>add.assoc\<close>*) (*"?z1.1 + (?z2.1 + ?z3.1) = ?z1.1 + ?z2.1 + ?z3.1"*)];
neuper@37950
   904
wneuper@59523
   905
val expand_poly =
walther@59851
   906
  Rule_Def.Repeat{id = "expand_poly", preconds = [], 
Walther@60509
   907
      rew_ord = ("dummy_ord", Rewrite_Ord.function_empty),
walther@60320
   908
      erls = powers_erls, srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@60358
   909
      rules = [
walther@60358
   910
        \<^rule_thm>\<open>distrib_right\<close>, (*"(z1.0 + z2.0) * w = z1.0 * w + z2.0 * w"*)
walther@60358
   911
        \<^rule_thm>\<open>distrib_left\<close>, (*"w * (z1.0 + z2.0) = w * z1.0 + w * z2.0"*)
neuper@37950
   912
walther@60358
   913
        \<^rule_thm>\<open>real_plus_binom_pow2\<close>, (*"(a + b) \<up> 2 = a \<up> 2 + 2*a*b + b \<up> 2"*)
walther@60358
   914
        \<^rule_thm>\<open>real_minus_binom_pow2_p\<close>, (*"(a - b) \<up> 2 = a \<up> 2 + -2*a*b + b \<up> 2"*)
walther@60358
   915
        \<^rule_thm>\<open>real_plus_minus_binom1_p\<close>, (*"(a + b)*(a - b) = a \<up> 2 + -1*b \<up> 2"*)
walther@60358
   916
        \<^rule_thm>\<open>real_plus_minus_binom2_p\<close>, (*"(a - b)*(a + b) = a \<up> 2 + -1*b \<up> 2"*)
neuper@37950
   917
walther@60358
   918
        \<^rule_thm>\<open>minus_minus\<close> (*"- (- ?z) = ?z"*),
walther@60358
   919
        \<^rule_thm>\<open>real_diff_minus\<close> (*"a - b = a + -1 * b"*),
walther@60385
   920
        \<^rule_thm>\<open>real_mult_minus1_sym\<close> (*"\<not>(z is_num) ==> - (z::real) = -1 * z"*)
neuper@37950
   921
walther@60358
   922
	       (*\<^rule_thm>\<open>real_minus_add_distrib\<close>,*) (*"- (?x + ?y) = - ?x + - ?y"*)
walther@60358
   923
	       (*\<^rule_thm>\<open>real_diff_plus\<close>*) (*"a - b = a + -b"*)],
walther@60358
   924
      scr = Rule.Empty_Prog};
neuper@37950
   925
neuper@37950
   926
val simplify_power = 
walther@59851
   927
  Rule_Def.Repeat{id = "simplify_power", preconds = [], 
Walther@60509
   928
      rew_ord = ("dummy_ord", Rewrite_Ord.function_empty),
walther@59852
   929
      erls = Rule_Set.empty, srls = Rule_Set.Empty,
neuper@42451
   930
      calc = [], errpatts = [],
walther@60358
   931
      rules = [
walther@60358
   932
        \<^rule_thm>\<open>realpow_multI\<close>, (*"(r * s) \<up> n = r \<up> n * s \<up> n"*)
walther@60358
   933
walther@60358
   934
        \<^rule_thm_sym>\<open>realpow_twoI\<close>, (*"r1 * r1 = r1 \<up> 2"*)
walther@60358
   935
        \<^rule_thm>\<open>realpow_plus_1\<close>, (*"r * r \<up> n = r \<up> (n + 1)"*)
walther@60358
   936
        \<^rule_thm>\<open>realpow_pow\<close>, (*"(a \<up> b) \<up> c = a \<up> (b * c)"*)
walther@60358
   937
        \<^rule_thm_sym>\<open>realpow_addI\<close>, (*"r \<up> n * r \<up> m = r \<up> (n + m)"*)
walther@60358
   938
        \<^rule_thm>\<open>realpow_oneI\<close>, (*"r \<up> 1 = r"*)
walther@60358
   939
        \<^rule_thm>\<open>realpow_eq_oneI\<close> (*"1 \<up> n = 1"*)],
walther@60358
   940
      scr = Rule.Empty_Prog};
neuper@42451
   941
neuper@37950
   942
val collect_numerals = 
walther@59851
   943
  Rule_Def.Repeat{id = "collect_numerals", preconds = [], 
Walther@60509
   944
      rew_ord = ("dummy_ord", Rewrite_Ord.function_empty),
walther@59851
   945
      erls = Atools_erls(*erls3.4.03*),srls = Rule_Set.Empty,
wenzelm@60309
   946
      calc = [("PLUS"  , (\<^const_name>\<open>plus\<close>, eval_binop "#add_")), 
wenzelm@60309
   947
	      ("TIMES" , (\<^const_name>\<open>times\<close>, eval_binop "#mult_")),
wenzelm@60405
   948
	      ("POWER", (\<^const_name>\<open>realpow\<close>, eval_binop "#power_"))
neuper@42451
   949
	      ], errpatts = [],
walther@60358
   950
      rules =[
walther@60385
   951
        \<^rule_thm>\<open>real_num_collect\<close>, (*"[| l is_num; m is_num |]==>l * n + m * n = (l + m) * n"*)
walther@60358
   952
	      \<^rule_thm>\<open>real_num_collect_assoc\<close>,
walther@60385
   953
          (*"[| l is_num; m is_num |] ==> l * n + (m * n + k) =  (l + m) * n + k"*)
walther@60385
   954
	      \<^rule_thm>\<open>real_one_collect\<close>,	 (*"m is_num ==> n + m * n = (1 + m) * n"*)
walther@60385
   955
	      \<^rule_thm>\<open>real_one_collect_assoc\<close>, (*"m is_num ==> k + (n + m * n) = k + (1 + m) * n"*)
walther@60358
   956
	      \<^rule_eval>\<open>plus\<close> (eval_binop "#add_"), 
walther@60358
   957
	      \<^rule_eval>\<open>times\<close> (eval_binop "#mult_"),
wenzelm@60405
   958
	      \<^rule_eval>\<open>realpow\<close> (eval_binop "#power_")],
walther@60358
   959
      scr = Rule.Empty_Prog};
neuper@37950
   960
val reduce_012 = 
walther@59851
   961
  Rule_Def.Repeat{id = "reduce_012", preconds = [], 
Walther@60509
   962
      rew_ord = ("dummy_ord", Rewrite_Ord.function_empty),
walther@60344
   963
      erls = Rule_Set.append_rules "erls_in_reduce_012" Rule_Set.empty [
walther@60344
   964
        \<^rule_eval>\<open>is_num\<close> (Prog_Expr.eval_is_num "#is_num_"),
walther@60344
   965
        \<^rule_thm>\<open>not_false\<close>,
walther@60344
   966
        \<^rule_thm>\<open>not_true\<close>],
walther@60344
   967
      srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@60358
   968
      rules = [
walther@60358
   969
        \<^rule_thm>\<open>mult_1_left\<close>, (*"1 * z = z"*)
walther@60358
   970
	       (*\<^rule_thm>\<open>real_mult_minus1\<close>,14.3.03*) (*"-1 * z = - z"*)
walther@60344
   971
	       \<^rule_thm>\<open>real_minus_mult_left\<close>, (*"\<not> ((a::real) is_num) ==> (- a) * b = - (a * b)"*)
walther@60358
   972
	       (*\<^rule_thm>\<open>real_minus_mult_cancel\<close>, (*"- ?x * - ?y = ?x * ?y"*)---*)
walther@60358
   973
	       \<^rule_thm>\<open>mult_zero_left\<close>, (*"0 * z = 0"*)
walther@60344
   974
	       \<^rule_thm>\<open>add_0_left\<close>, (*"0 + z = z"*)
walther@60344
   975
	       \<^rule_thm>\<open>add_0_right\<close>, (*"a + - a = 0"*)
walther@60358
   976
	       \<^rule_thm>\<open>right_minus\<close>, (*"?z + - ?z = 0"*)
walther@60358
   977
	       \<^rule_thm_sym>\<open>real_mult_2\<close>,	 (*"z1 + z1 = 2 * z1"*)
walther@60358
   978
	       \<^rule_thm>\<open>real_mult_2_assoc\<close> (*"z1 + (z1 + k) = 2 * z1 + k"*)],
walther@60358
   979
      scr = Rule.Empty_Prog};
neuper@52139
   980
neuper@37950
   981
val discard_parentheses = 
walther@59852
   982
    Rule_Set.append_rules "discard_parentheses" Rule_Set.empty 
wenzelm@60296
   983
	       [\<^rule_thm_sym>\<open>mult.assoc\<close>,	\<^rule_thm_sym>\<open>add.assoc\<close>];
wneuper@59523
   984
\<close>
neuper@37950
   985
wneuper@59523
   986
subsubsection \<open>hard-coded AC rewriting\<close>
wneuper@59523
   987
ML \<open>
wneuper@59523
   988
(*MG.0401: termorders for multivariate polys dropped due to principal problems:
wneuper@59523
   989
  (total-degree-)ordering of monoms NOT possible with size_of_term GIVEN*)
wneuper@59523
   990
val order_add_mult = 
walther@59851
   991
  Rule_Def.Repeat{id = "order_add_mult", preconds = [], 
Walther@60509
   992
      rew_ord = ("ord_make_polynomial", ord_make_polynomial false \<^theory>),
walther@59852
   993
      erls = Rule_Set.empty,srls = Rule_Set.Empty,
neuper@42451
   994
      calc = [], errpatts = [],
walther@60358
   995
      rules = [
walther@60358
   996
        \<^rule_thm>\<open>mult.commute\<close>, (* z * w = w * z *)
walther@60358
   997
	       \<^rule_thm>\<open>real_mult_left_commute\<close>, (*z1.0 * (z2.0 * z3.0) = z2.0 * (z1.0 * z3.0)*)
walther@60358
   998
	       \<^rule_thm>\<open>mult.assoc\<close>, (*z1.0 * z2.0 * z3.0 = z1.0 * (z2.0 * z3.0)*)
walther@60358
   999
	       \<^rule_thm>\<open>add.commute\<close>,	 (*z + w = w + z*)
walther@60358
  1000
	       \<^rule_thm>\<open>add.left_commute\<close>, (*x + (y + z) = y + (x + z)*)
walther@60358
  1001
	       \<^rule_thm>\<open>add.assoc\<close> (*z1.0 + z2.0 + z3.0 = z1.0 + (z2.0 + z3.0)*)],
walther@60358
  1002
      scr = Rule.Empty_Prog};
wneuper@59523
  1003
(*MG.0401: termorders for multivariate polys dropped due to principal problems:
wneuper@59523
  1004
  (total-degree-)ordering of monoms NOT possible with size_of_term GIVEN*)
wneuper@59523
  1005
val order_mult = 
walther@59851
  1006
  Rule_Def.Repeat{id = "order_mult", preconds = [], 
wenzelm@60291
  1007
      rew_ord = ("ord_make_polynomial",ord_make_polynomial false \<^theory>),
walther@59852
  1008
      erls = Rule_Set.empty,srls = Rule_Set.Empty,
wneuper@59523
  1009
      calc = [], errpatts = [],
walther@60358
  1010
      rules = [
walther@60358
  1011
        \<^rule_thm>\<open>mult.commute\<close>, (* z * w = w * z *)
walther@60358
  1012
	       \<^rule_thm>\<open>real_mult_left_commute\<close>, (*z1.0 * (z2.0 * z3.0) = z2.0 * (z1.0 * z3.0)*)
walther@60358
  1013
	       \<^rule_thm>\<open>mult.assoc\<close>	 (*z1.0 * z2.0 * z3.0 = z1.0 * (z2.0 * z3.0)*)],
walther@60358
  1014
      scr = Rule.Empty_Prog};
wneuper@59472
  1015
\<close>
wneuper@59472
  1016
ML \<open>
wneuper@59416
  1017
fun attach_form (_: Rule.rule list list) (_: term) (_: term) = (*still missing*)
wneuper@59416
  1018
    []:(Rule.rule * (term * term list)) list;
walther@59850
  1019
fun init_state (_: term) = Rule_Set.e_rrlsstate;
wneuper@59416
  1020
fun locate_rule (_: Rule.rule list list) (_: term) (_: Rule.rule) =
wneuper@59416
  1021
    ([]:(Rule.rule * (term * term list)) list);
wneuper@59416
  1022
fun next_rule (_: Rule.rule list list) (_: term) = (NONE: Rule.rule option);
wneuper@59406
  1023
fun normal_form t = SOME (sort_variables t, []: term list);
neuper@37950
  1024
neuper@37950
  1025
val order_mult_ =
walther@59850
  1026
    Rule_Set.Rrls {id = "order_mult_", 
neuper@37950
  1027
	  prepat = 
neuper@38036
  1028
          (* ?p matched with the current term gives an environment,
neuper@38037
  1029
             which evaluates (the instantiated) "?p is_multUnordered" to true *)
wenzelm@60291
  1030
	  [([TermC.parse_patt \<^theory> "?p is_multUnordered"], 
wenzelm@60291
  1031
             TermC.parse_patt \<^theory> "?p :: real")],
Walther@60509
  1032
	  rew_ord = ("dummy_ord", Rewrite_Ord.function_empty),
walther@59852
  1033
	  erls = Rule_Set.append_rules "Rule_Set.empty-is_multUnordered" Rule_Set.empty
wenzelm@60294
  1034
			    [\<^rule_eval>\<open>is_multUnordered\<close> (eval_is_multUnordered "")],
wenzelm@60309
  1035
	  calc = [("PLUS"  , (\<^const_name>\<open>plus\<close>, eval_binop "#add_")),
wenzelm@60309
  1036
		  ("TIMES" , (\<^const_name>\<open>times\<close>, eval_binop "#mult_")),
wenzelm@60309
  1037
		  ("DIVIDE", (\<^const_name>\<open>divide\<close>, Prog_Expr.eval_cancel "#divide_e")),
wenzelm@60405
  1038
		  ("POWER" , (\<^const_name>\<open>realpow\<close>, eval_binop "#power_"))],
wneuper@59406
  1039
    errpatts = [],
wneuper@59416
  1040
	  scr = Rule.Rfuns {init_state  = init_state,
neuper@37950
  1041
		     normal_form = normal_form,
neuper@37950
  1042
		     locate_rule = locate_rule,
neuper@37950
  1043
		     next_rule   = next_rule,
neuper@37950
  1044
		     attach_form = attach_form}};
neuper@37950
  1045
val order_mult_rls_ = 
walther@59851
  1046
  Rule_Def.Repeat {id = "order_mult_rls_", preconds = [], 
Walther@60509
  1047
    rew_ord = ("dummy_ord", Rewrite_Ord.function_empty),
walther@60358
  1048
    erls = Rule_Set.empty,srls = Rule_Set.Empty,
walther@60358
  1049
    calc = [], errpatts = [],
walther@60358
  1050
    rules = [
walther@60358
  1051
      Rule.Rls_ order_mult_],
walther@60358
  1052
    scr = Rule.Empty_Prog};
neuper@37950
  1053
wneuper@59523
  1054
\<close> ML \<open>
neuper@37950
  1055
wneuper@59416
  1056
fun attach_form (_: Rule.rule list list) (_: term) (_: term) = (*still missing*)
wneuper@59416
  1057
    []: (Rule.rule * (term * term list)) list;
walther@59850
  1058
fun init_state (_: term) = Rule_Set.e_rrlsstate;
wneuper@59416
  1059
fun locate_rule (_: Rule.rule list list) (_: term) (_: Rule.rule) =
wneuper@59416
  1060
    ([]: (Rule.rule * (term * term list)) list);
wneuper@59416
  1061
fun next_rule (_: Rule.rule list list) (_: term) = (NONE: Rule.rule option);
wneuper@59406
  1062
fun normal_form t = SOME (sort_monoms t,[]: term list);
wneuper@59472
  1063
\<close> ML \<open>
neuper@37950
  1064
val order_add_ =
walther@60358
  1065
  Rule_Set.Rrls {id = "order_add_", 
walther@60358
  1066
	  prepat = [(*WN.18.6.03 Preconditions und Pattern,
walther@59850
  1067
		    die beide passen muessen, damit das Rule_Set.Rrls angewandt wird*)
walther@60358
  1068
	    ([TermC.parse_patt @{theory} "?p is_addUnordered"], 
walther@60358
  1069
	      TermC.parse_patt @{theory} "?p :: real" 
neuper@37950
  1070
	    (*WN.18.6.03 also KEIN pattern, dieses erzeugt nur das Environment 
neuper@37950
  1071
	      fuer die Evaluation der Precondition "p is_addUnordered"*))],
Walther@60509
  1072
	  rew_ord = ("dummy_ord", Rewrite_Ord.function_empty),
walther@59852
  1073
	  erls = Rule_Set.append_rules "Rule_Set.empty-is_addUnordered" Rule_Set.empty(*MG: poly_erls*)
walther@60358
  1074
			[\<^rule_eval>\<open>is_addUnordered\<close> (eval_is_addUnordered "")],
walther@60358
  1075
	  calc = [
walther@60358
  1076
      ("PLUS"  ,(\<^const_name>\<open>plus\<close>, eval_binop "#add_")),
wenzelm@60309
  1077
		  ("TIMES" ,(\<^const_name>\<open>times\<close>, eval_binop "#mult_")),
wenzelm@60309
  1078
		  ("DIVIDE",(\<^const_name>\<open>divide\<close>, Prog_Expr.eval_cancel "#divide_e")),
wenzelm@60405
  1079
		  ("POWER" ,(\<^const_name>\<open>realpow\<close>  , eval_binop "#power_"))],
neuper@42451
  1080
	  errpatts = [],
walther@60358
  1081
	  scr = Rule.Rfuns {
walther@60358
  1082
      init_state  = init_state,
walther@60358
  1083
		  normal_form = normal_form,
walther@60358
  1084
		  locate_rule = locate_rule,
walther@60358
  1085
		  next_rule   = next_rule,
walther@60358
  1086
		  attach_form = attach_form}};
neuper@37950
  1087
wneuper@59406
  1088
val order_add_rls_ =
walther@59851
  1089
  Rule_Def.Repeat {id = "order_add_rls_", preconds = [], 
Walther@60509
  1090
    rew_ord = ("dummy_ord", Rewrite_Ord.function_empty),
walther@60358
  1091
    erls = Rule_Set.empty, srls = Rule_Set.Empty,
walther@60358
  1092
    calc = [], errpatts = [],
walther@60358
  1093
    rules = [
walther@60358
  1094
      Rule.Rls_ order_add_],
walther@60358
  1095
    scr = Rule.Empty_Prog};
wneuper@59472
  1096
\<close>
neuper@37950
  1097
wneuper@59472
  1098
text \<open>rule-set make_polynomial also named norm_Poly:
neuper@42398
  1099
  Rewrite order has not been implemented properly; the order is better in 
neuper@42398
  1100
  make_polynomial_in (coded in SML).
neuper@42398
  1101
  Notes on state of development:
neuper@42398
  1102
  \# surprise 2006: test --- norm_Poly NOT COMPLETE ---
neuper@42398
  1103
  \# migration Isabelle2002 --> 2011 weakened the rule set, see test
walther@59962
  1104
  --- Matthias Goldgruber 2003 rewrite orders ---, raise ERROR "ord_make_polynomial_in #16b"
wneuper@59472
  1105
\<close>
wneuper@59472
  1106
ML \<open>
neuper@37950
  1107
(*. see MG-DA.p.52ff .*)
neuper@37950
  1108
val make_polynomial(*MG.03, overwrites version from above, 
neuper@37950
  1109
    previously 'make_polynomial_'*) =
walther@59878
  1110
  Rule_Set.Sequence {id = "make_polynomial", preconds = []:term list, 
Walther@60509
  1111
    rew_ord = ("dummy_ord", Rewrite_Ord.function_empty),
walther@60358
  1112
    erls = Atools_erls, srls = Rule_Set.Empty,calc = [], errpatts = [],
walther@60358
  1113
    rules = [
walther@60358
  1114
       Rule.Rls_ discard_minus,
walther@60358
  1115
	     Rule.Rls_ expand_poly_,
walther@60358
  1116
	     \<^rule_eval>\<open>times\<close> (eval_binop "#mult_"),
walther@60358
  1117
	     Rule.Rls_ order_mult_rls_,
walther@60358
  1118
	     Rule.Rls_ simplify_power_, 
walther@60358
  1119
	     Rule.Rls_ calc_add_mult_pow_, 
walther@60358
  1120
	     Rule.Rls_ reduce_012_mult_,
walther@60358
  1121
	     Rule.Rls_ order_add_rls_,
walther@60358
  1122
	     Rule.Rls_ collect_numerals_, 
walther@60358
  1123
	     Rule.Rls_ reduce_012_,
walther@60358
  1124
	     Rule.Rls_ discard_parentheses1],
walther@60358
  1125
    scr = Rule.Empty_Prog};
wneuper@59472
  1126
\<close>
wneuper@59472
  1127
ML \<open>
neuper@37950
  1128
val norm_Poly(*=make_polynomial*) = 
walther@59878
  1129
  Rule_Set.Sequence {id = "norm_Poly", preconds = []:term list, 
Walther@60509
  1130
    rew_ord = ("dummy_ord", Rewrite_Ord.function_empty),
walther@60358
  1131
    erls = Atools_erls, srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@60358
  1132
    rules = [
walther@60358
  1133
       Rule.Rls_ discard_minus,
walther@60358
  1134
	     Rule.Rls_ expand_poly_,
walther@60358
  1135
	     \<^rule_eval>\<open>times\<close> (eval_binop "#mult_"),
walther@60358
  1136
	     Rule.Rls_ order_mult_rls_,
walther@60358
  1137
	     Rule.Rls_ simplify_power_, 
walther@60358
  1138
	     Rule.Rls_ calc_add_mult_pow_, 
walther@60358
  1139
	     Rule.Rls_ reduce_012_mult_,
walther@60358
  1140
	     Rule.Rls_ order_add_rls_,
walther@60358
  1141
	     Rule.Rls_ collect_numerals_, 
walther@60358
  1142
	     Rule.Rls_ reduce_012_,
walther@60358
  1143
	     Rule.Rls_ discard_parentheses1],
walther@60358
  1144
    scr = Rule.Empty_Prog};
wneuper@59472
  1145
\<close>
wneuper@59472
  1146
ML \<open>
wneuper@59416
  1147
(* MG:03 Like make_polynomial_ but without Rule.Rls_ discard_parentheses1 
neuper@37950
  1148
   and expand_poly_rat_ instead of expand_poly_, see MG-DA.p.56ff*)
neuper@37950
  1149
(* MG necessary  for termination of norm_Rational(*_mg*) in Rational.ML*)
neuper@37950
  1150
val make_rat_poly_with_parentheses =
walther@59878
  1151
  Rule_Set.Sequence{id = "make_rat_poly_with_parentheses", preconds = []:term list, 
Walther@60509
  1152
    rew_ord = ("dummy_ord", Rewrite_Ord.function_empty),
walther@60358
  1153
    erls = Atools_erls, srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@60358
  1154
    rules = [
walther@60358
  1155
      Rule.Rls_ discard_minus,
walther@60358
  1156
	    Rule.Rls_ expand_poly_rat_,(*ignors rationals*)
walther@60358
  1157
	    \<^rule_eval>\<open>times\<close> (eval_binop "#mult_"),
walther@60358
  1158
	    Rule.Rls_ order_mult_rls_,
walther@60358
  1159
	    Rule.Rls_ simplify_power_, 
walther@60358
  1160
	    Rule.Rls_ calc_add_mult_pow_, 
walther@60358
  1161
	    Rule.Rls_ reduce_012_mult_,
walther@60358
  1162
	    Rule.Rls_ order_add_rls_,
walther@60358
  1163
	    Rule.Rls_ collect_numerals_, 
walther@60358
  1164
	    Rule.Rls_ reduce_012_
walther@60358
  1165
	    (*Rule.Rls_ discard_parentheses1 *)],
walther@60358
  1166
    scr = Rule.Empty_Prog};
wneuper@59472
  1167
\<close>
wneuper@59472
  1168
ML \<open>
neuper@37950
  1169
(*.a minimal ruleset for reverse rewriting of factions [2];
neuper@37950
  1170
   compare expand_binoms.*)
neuper@37950
  1171
val rev_rew_p = 
walther@60358
  1172
  Rule_Set.Sequence{id = "rev_rew_p", preconds = [], rew_ord = ("termlessI",termlessI),
walther@59851
  1173
    erls = Atools_erls, srls = Rule_Set.Empty,
wenzelm@60309
  1174
    calc = [(*("PLUS"  , (\<^const_name>\<open>plus\<close>, eval_binop "#add_")), 
wenzelm@60309
  1175
	    ("TIMES" , (\<^const_name>\<open>times\<close>, eval_binop "#mult_")),
wenzelm@60405
  1176
	    ("POWER", (\<^const_name>\<open>realpow\<close>, eval_binop "#power_"))*)
neuper@42451
  1177
	    ], errpatts = [],
walther@60358
  1178
    rules = [
walther@60358
  1179
       \<^rule_thm>\<open>real_plus_binom_times\<close>, (*"(a + b)*(a + b) = a ^ 2 + 2 * a * b + b ^ 2*)
walther@60358
  1180
	     \<^rule_thm>\<open>real_plus_binom_times1\<close>, (*"(a +  1*b)*(a + -1*b) = a \<up> 2 + -1*b \<up> 2"*)
walther@60358
  1181
	     \<^rule_thm>\<open>real_plus_binom_times2\<close>, (*"(a + -1*b)*(a +  1*b) = a \<up> 2 + -1*b \<up> 2"*)
neuper@37950
  1182
wenzelm@60297
  1183
	     \<^rule_thm>\<open>mult_1_left\<close>,(*"1 * z = z"*)
neuper@37950
  1184
walther@60358
  1185
       \<^rule_thm>\<open>distrib_right\<close>, (*"(z1.0 + z2.0) * w = z1.0 * w + z2.0 * w"*)
walther@60358
  1186
	     \<^rule_thm>\<open>distrib_left\<close>, (*"w * (z1.0 + z2.0) = w * z1.0 + w * z2.0"*)
neuper@37950
  1187
	       
walther@60358
  1188
	     \<^rule_thm>\<open>mult.assoc\<close>, (*"?z1.1 * ?z2.1 * ?z3. =1 ?z1.1 * (?z2.1 * ?z3.1)"*)
wneuper@59416
  1189
	     Rule.Rls_ order_mult_rls_,
wneuper@59416
  1190
	     (*Rule.Rls_ order_add_rls_,*)
neuper@37950
  1191
wenzelm@60294
  1192
	     \<^rule_eval>\<open>plus\<close> (eval_binop "#add_"), 
wenzelm@60294
  1193
	     \<^rule_eval>\<open>times\<close> (eval_binop "#mult_"),
wenzelm@60405
  1194
	     \<^rule_eval>\<open>realpow\<close> (eval_binop "#power_"),
neuper@37950
  1195
	     
walther@60358
  1196
	     \<^rule_thm_sym>\<open>realpow_twoI\<close>, (*"r1 * r1 = r1 \<up> 2"*)
walther@60358
  1197
	     \<^rule_thm_sym>\<open>real_mult_2\<close>, (*"z1 + z1 = 2 * z1"*)
walther@60358
  1198
	     \<^rule_thm>\<open>real_mult_2_assoc\<close>, (*"z1 + (z1 + k) = 2 * z1 + k"*)
neuper@37950
  1199
walther@60385
  1200
	     \<^rule_thm>\<open>real_num_collect\<close>,  (*"[| l is_num; m is_num |]==>l * n + m * n = (l + m) * n"*)
walther@60385
  1201
	     \<^rule_thm>\<open>real_num_collect_assoc\<close>, (*"[| l is_num; m is_num |] ==> l * n + (m * n + k) =  (l + m) * n + k"*)
walther@60385
  1202
	     \<^rule_thm>\<open>real_one_collect\<close>, (*"m is_num ==> n + m * n = (1 + m) * n"*)
walther@60385
  1203
	     \<^rule_thm>\<open>real_one_collect_assoc\<close>, (*"m is_num ==> k + (n + m * n) = k + (1 + m) * n"*)
neuper@37950
  1204
walther@60358
  1205
	     \<^rule_thm>\<open>realpow_multI\<close>, (*"(r * s) \<up> n = r \<up> n * s \<up> n"*)
neuper@37950
  1206
wenzelm@60294
  1207
	     \<^rule_eval>\<open>plus\<close> (eval_binop "#add_"),
wenzelm@60294
  1208
	     \<^rule_eval>\<open>times\<close> (eval_binop "#mult_"),
wenzelm@60405
  1209
	     \<^rule_eval>\<open>realpow\<close> (eval_binop "#power_"),
neuper@37950
  1210
walther@60358
  1211
	     \<^rule_thm>\<open>mult_1_left\<close>, (*"1 * z = z"*)
walther@60358
  1212
	     \<^rule_thm>\<open>mult_zero_left\<close>, (*"0 * z = 0"*)
walther@60358
  1213
	     \<^rule_thm>\<open>add_0_left\<close> (*0 + z = z*)
neuper@37950
  1214
wneuper@59416
  1215
	     (*Rule.Rls_ order_add_rls_*)
neuper@37950
  1216
	     ],
walther@59878
  1217
    scr = Rule.Empty_Prog};      
wneuper@59472
  1218
\<close>
neuper@52125
  1219
wneuper@59523
  1220
subsection \<open>rule-sets with explicit program for intermediate steps\<close>
wneuper@59523
  1221
partial_function (tailrec) expand_binoms_2 :: "real \<Rightarrow> real"
wneuper@59523
  1222
  where
walther@59635
  1223
"expand_binoms_2 term = (
walther@59635
  1224
  Repeat (
walther@59637
  1225
    (Try (Repeat (Rewrite ''real_plus_binom_pow2''))) #>
walther@59637
  1226
    (Try (Repeat (Rewrite ''real_plus_binom_times''))) #>
walther@59637
  1227
    (Try (Repeat (Rewrite ''real_minus_binom_pow2''))) #>
walther@59637
  1228
    (Try (Repeat (Rewrite ''real_minus_binom_times''))) #>
walther@59637
  1229
    (Try (Repeat (Rewrite ''real_plus_minus_binom1''))) #>
walther@59637
  1230
    (Try (Repeat (Rewrite ''real_plus_minus_binom2''))) #>
walther@59635
  1231
   
walther@59637
  1232
    (Try (Repeat (Rewrite ''mult_1_left''))) #>
walther@59637
  1233
    (Try (Repeat (Rewrite ''mult_zero_left''))) #>
walther@59637
  1234
    (Try (Repeat (Rewrite ''add_0_left''))) #>
walther@59635
  1235
   
walther@59637
  1236
    (Try (Repeat (Calculate ''PLUS''))) #>
walther@59637
  1237
    (Try (Repeat (Calculate ''TIMES''))) #>
walther@59637
  1238
    (Try (Repeat (Calculate ''POWER''))) #>
walther@59635
  1239
   
walther@59637
  1240
    (Try (Repeat (Rewrite ''sym_realpow_twoI''))) #>
walther@59637
  1241
    (Try (Repeat (Rewrite ''realpow_plus_1''))) #>
walther@59637
  1242
    (Try (Repeat (Rewrite ''sym_real_mult_2''))) #>
walther@59637
  1243
    (Try (Repeat (Rewrite ''real_mult_2_assoc''))) #>
walther@59635
  1244
   
walther@59637
  1245
    (Try (Repeat (Rewrite ''real_num_collect''))) #>
walther@59637
  1246
    (Try (Repeat (Rewrite ''real_num_collect_assoc''))) #>
walther@59635
  1247
   
walther@59637
  1248
    (Try (Repeat (Rewrite ''real_one_collect''))) #>
walther@59637
  1249
    (Try (Repeat (Rewrite ''real_one_collect_assoc''))) #>
walther@59635
  1250
   
walther@59637
  1251
    (Try (Repeat (Calculate ''PLUS''))) #>
walther@59637
  1252
    (Try (Repeat (Calculate ''TIMES''))) #>
walther@59635
  1253
    (Try (Repeat (Calculate ''POWER''))))
walther@59635
  1254
  term)"
wneuper@59523
  1255
ML \<open>
wneuper@59523
  1256
val expand_binoms = 
walther@59851
  1257
  Rule_Def.Repeat{id = "expand_binoms", preconds = [], rew_ord = ("termlessI",termlessI),
walther@60358
  1258
    erls = Atools_erls, srls = Rule_Set.Empty,
walther@60358
  1259
    calc = [("PLUS"  , (\<^const_name>\<open>plus\<close>, eval_binop "#add_")), 
walther@60358
  1260
	    ("TIMES" , (\<^const_name>\<open>times\<close>, eval_binop "#mult_")),
wenzelm@60405
  1261
	    ("POWER", (\<^const_name>\<open>realpow\<close>, eval_binop "#power_"))
walther@60358
  1262
	    ], errpatts = [],
walther@60358
  1263
    rules = [
walther@60358
  1264
       \<^rule_thm>\<open>real_plus_binom_pow2\<close>, (*"(a + b) \<up> 2 = a \<up> 2 + 2 * a * b + b \<up> 2"*)
walther@60358
  1265
	     \<^rule_thm>\<open>real_plus_binom_times\<close>, (*"(a + b)*(a + b) = ...*)
walther@60358
  1266
	     \<^rule_thm>\<open>real_minus_binom_pow2\<close>,  (*"(a - b) \<up> 2 = a \<up> 2 - 2 * a * b + b \<up> 2"*)
walther@60358
  1267
	     \<^rule_thm>\<open>real_minus_binom_times\<close>, (*"(a - b)*(a - b) = ...*)
walther@60358
  1268
	     \<^rule_thm>\<open>real_plus_minus_binom1\<close>, (*"(a + b) * (a - b) = a \<up> 2 - b \<up> 2"*)
walther@60358
  1269
	     \<^rule_thm>\<open>real_plus_minus_binom2\<close>, (*"(a - b) * (a + b) = a \<up> 2 - b \<up> 2"*)
walther@60358
  1270
	     (*RL 020915*)
walther@60358
  1271
	     \<^rule_thm>\<open>real_pp_binom_times\<close>, (*(a + b)*(c + d) = a*c + a*d + b*c + b*d*)
walther@60384
  1272
	     \<^rule_thm>\<open>real_pm_binom_times\<close>, (*(a + b)*(c - d) = a*c - a*d + b*c - b*d*)
walther@60384
  1273
	     \<^rule_thm>\<open>real_mp_binom_times\<close>, (*(a - b)*(c + d) = a*c + a*d - b*c - b*d*)
walther@60384
  1274
	     \<^rule_thm>\<open>real_mm_binom_times\<close>, (*(a - b)*(c - d) = a*c - a*d - b*c + b*d*)
walther@60358
  1275
	     \<^rule_thm>\<open>realpow_multI\<close>, (*(a*b) \<up> n = a \<up> n * b \<up> n*)
walther@60358
  1276
	     \<^rule_thm>\<open>real_plus_binom_pow3\<close>, (* (a + b) \<up> 3 = a \<up> 3 + 3*a \<up> 2*b + 3*a*b \<up> 2 + b \<up> 3 *)
walther@60358
  1277
	     \<^rule_thm>\<open>real_minus_binom_pow3\<close>, (* (a - b) \<up> 3 = a \<up> 3 - 3*a \<up> 2*b + 3*a*b \<up> 2 - b \<up> 3 *)
walther@60358
  1278
       (*
walther@60358
  1279
       \<^rule_thm>\<open>distrib_right\<close>,	 (*"(z1.0 + z2.0) * w = z1.0 * w + z2.0 * w"*)
walther@60358
  1280
	     \<^rule_thm>\<open>distrib_left\<close>, (*"w * (z1.0 + z2.0) = w * z1.0 + w * z2.0"*)
walther@60358
  1281
	     \<^rule_thm>\<open>left_diff_distrib\<close>,	 (*"(z1.0 - z2.0) * w = z1.0 * w - z2.0 * w"*)
walther@60358
  1282
	     \<^rule_thm>\<open>right_diff_distrib\<close>, (*"w * (z1.0 - z2.0) = w * z1.0 - w * z2.0"*)
wneuper@59523
  1283
	      *)
walther@60358
  1284
	     \<^rule_thm>\<open>mult_1_left\<close>, (*"1 * z = z"*)
walther@60358
  1285
	     \<^rule_thm>\<open>mult_zero_left\<close>, (*"0 * z = 0"*)
walther@60358
  1286
	     \<^rule_thm>\<open>add_0_left\<close>, (*"0 + z = z"*)
walther@60358
  1287
    
walther@60358
  1288
	     \<^rule_eval>\<open>plus\<close> (eval_binop "#add_"), 
walther@60358
  1289
	     \<^rule_eval>\<open>times\<close> (eval_binop "#mult_"),
wenzelm@60405
  1290
	     \<^rule_eval>\<open>realpow\<close> (eval_binop "#power_"),
walther@60358
  1291
            (*\<^rule_thm>\<open>mult.commute\<close>,
walther@60358
  1292
		    (*AC-rewriting*)
walther@60358
  1293
	     \<^rule_thm>\<open>real_mult_left_commute\<close>,
walther@60358
  1294
	     \<^rule_thm>\<open>mult.assoc\<close>,
walther@60358
  1295
	     \<^rule_thm>\<open>add.commute\<close>,
walther@60358
  1296
	     \<^rule_thm>\<open>add.left_commute\<close>,
walther@60358
  1297
	     \<^rule_thm>\<open>add.assoc\<close>,
wneuper@59523
  1298
	      *)
walther@60358
  1299
	     \<^rule_thm_sym>\<open>realpow_twoI\<close>, (*"r1 * r1 = r1 \<up> 2"*)
walther@60358
  1300
	     \<^rule_thm>\<open>realpow_plus_1\<close>, (*"r * r \<up> n = r \<up> (n + 1)"*)
walther@60358
  1301
	     (*
walther@60358
  1302
       \<^rule_thm_sym>\<open>real_mult_2\<close>, (*"z1 + z1 = 2 * z1"*)*)
walther@60358
  1303
	     \<^rule_thm>\<open>real_mult_2_assoc\<close>,	 (*"z1 + (z1 + k) = 2 * z1 + k"*)
walther@60358
  1304
    
walther@60385
  1305
	     \<^rule_thm>\<open>real_num_collect\<close>, (*"[| l is_num; m is_num |] ==>l * n + m * n = (l + m) * n"*)
walther@60385
  1306
	     \<^rule_thm>\<open>real_num_collect_assoc\<close>, (*"[| l is_num; m is_num |] ==> l * n + (m * n + k) =  (l + m) * n + k"*)
walther@60385
  1307
	     \<^rule_thm>\<open>real_one_collect\<close>, (*"m is_num ==> n + m * n = (1 + m) * n"*)
walther@60385
  1308
	     \<^rule_thm>\<open>real_one_collect_assoc\<close>, (*"m is_num ==> k + (n + m * n) = k + (1 + m) * n"*)
walther@60358
  1309
    
walther@60358
  1310
	     \<^rule_eval>\<open>plus\<close> (eval_binop "#add_"), 
walther@60358
  1311
	     \<^rule_eval>\<open>times\<close> (eval_binop "#mult_"),
wenzelm@60405
  1312
	     \<^rule_eval>\<open>realpow\<close> (eval_binop "#power_")],
walther@60358
  1313
    scr = Rule.Prog (Program.prep_program @{thm expand_binoms_2.simps})};      
wneuper@59523
  1314
\<close>
wneuper@59523
  1315
walther@59887
  1316
subsection \<open>add to Know_Store\<close>
wneuper@59523
  1317
subsubsection \<open>rule-sets\<close>
walther@59618
  1318
ML \<open>val prep_rls' = Auto_Prog.prep_rls @{theory}\<close>
s1210629013@55444
  1319
wenzelm@60289
  1320
rule_set_knowledge
walther@60387
  1321
  norm_Poly = \<open>prep_rls' norm_Poly\<close> and
walther@60387
  1322
  Poly_erls = \<open>prep_rls' Poly_erls\<close> and
wenzelm@60286
  1323
  expand = \<open>prep_rls' expand\<close> and
wenzelm@60286
  1324
  expand_poly = \<open>prep_rls' expand_poly\<close> and
wenzelm@60286
  1325
  simplify_power = \<open>prep_rls' simplify_power\<close> and
neuper@52125
  1326
wenzelm@60286
  1327
  order_add_mult = \<open>prep_rls' order_add_mult\<close> and
wenzelm@60286
  1328
  collect_numerals = \<open>prep_rls' collect_numerals\<close> and
wenzelm@60286
  1329
  collect_numerals_= \<open>prep_rls' collect_numerals_\<close> and
wenzelm@60286
  1330
  reduce_012 = \<open>prep_rls' reduce_012\<close> and
wenzelm@60286
  1331
  discard_parentheses = \<open>prep_rls' discard_parentheses\<close> and
neuper@52125
  1332
 
wenzelm@60286
  1333
  make_polynomial = \<open>prep_rls' make_polynomial\<close> and
wenzelm@60286
  1334
  expand_binoms = \<open>prep_rls' expand_binoms\<close> and
wenzelm@60286
  1335
  rev_rew_p = \<open>prep_rls' rev_rew_p\<close> and
wenzelm@60286
  1336
  discard_minus = \<open>prep_rls' discard_minus\<close> and
wenzelm@60286
  1337
  expand_poly_ = \<open>prep_rls' expand_poly_\<close> and
neuper@52125
  1338
 
wenzelm@60286
  1339
  expand_poly_rat_ = \<open>prep_rls' expand_poly_rat_\<close> and
wenzelm@60286
  1340
  simplify_power_ = \<open>prep_rls' simplify_power_\<close> and
wenzelm@60286
  1341
  calc_add_mult_pow_ = \<open>prep_rls' calc_add_mult_pow_\<close> and
wenzelm@60286
  1342
  reduce_012_mult_ = \<open>prep_rls' reduce_012_mult_\<close> and
wenzelm@60286
  1343
  reduce_012_ = \<open>prep_rls' reduce_012_\<close> and
neuper@52125
  1344
 
wenzelm@60286
  1345
  discard_parentheses1 = \<open>prep_rls' discard_parentheses1\<close> and
wenzelm@60286
  1346
  order_mult_rls_ = \<open>prep_rls' order_mult_rls_\<close> and
wenzelm@60286
  1347
  order_add_rls_ = \<open>prep_rls' order_add_rls_\<close> and
wenzelm@60286
  1348
  make_rat_poly_with_parentheses = \<open>prep_rls' make_rat_poly_with_parentheses\<close>
wneuper@59523
  1349
wneuper@59526
  1350
subsection \<open>problems\<close>
wenzelm@60306
  1351
wenzelm@60306
  1352
problem pbl_simp_poly : "polynomial/simplification" =
wenzelm@60306
  1353
  \<open>Rule_Set.append_rules "empty" Rule_Set.empty [(*for preds in where_*)
wenzelm@60306
  1354
    \<^rule_eval>\<open>is_polyexp\<close> (eval_is_polyexp "")]\<close>
Walther@60449
  1355
  Method_Ref: "simplification/for_polynomials"
wenzelm@60306
  1356
  CAS: "Simplify t_t"
wenzelm@60306
  1357
  Given: "Term t_t"
wenzelm@60306
  1358
  Where: "t_t is_polyexp"
wenzelm@60306
  1359
  Find: "normalform n_n"
wneuper@59429
  1360
wneuper@59526
  1361
subsection \<open>methods\<close>
wneuper@59545
  1362
wneuper@59429
  1363
partial_function (tailrec) simplify :: "real \<Rightarrow> real"
wneuper@59429
  1364
  where
walther@59635
  1365
"simplify term = ((Rewrite_Set ''norm_Poly'') term)"
wenzelm@60303
  1366
wenzelm@60303
  1367
method met_simp_poly : "simplification/for_polynomials" =
wenzelm@60303
  1368
  \<open>{rew_ord'="tless_true", rls' = Rule_Set.empty, calc = [], srls = Rule_Set.empty,
wenzelm@60303
  1369
    prls = Rule_Set.append_rules "simplification_for_polynomials_prls" Rule_Set.empty
wenzelm@60303
  1370
      [(*for preds in where_*) \<^rule_eval>\<open>is_polyexp\<close> (eval_is_polyexp"")],
wenzelm@60303
  1371
    crls = Rule_Set.empty, errpats = [], nrls = norm_Poly}\<close>
wenzelm@60303
  1372
  Program: simplify.simps
wenzelm@60303
  1373
  Given: "Term t_t"
wenzelm@60303
  1374
  Where: "t_t is_polyexp"
wenzelm@60303
  1375
  Find: "normalform n_n"
wneuper@59472
  1376
ML \<open>
wneuper@59472
  1377
\<close> ML \<open>
Walther@60509
  1378
\<close>
neuper@37906
  1379
end