src/Tools/isac/Knowledge/Poly.thy
author wneuper <walther.neuper@jku.at>
Tue, 13 Jul 2021 08:52:35 +0200
changeset 60320 02102eaa2021
parent 60318 e6e7a9b9ced7
child 60321 66086b5d1b6e
permissions -rw-r--r--
Test_Some.thy + rewrite.sml + poly.sml ok: real_mult_minus1_sym works with is_atom
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
neuper@37974
   112
(* real_plus_binom_pow:        "[| n is_const;  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
neuper@37974
   145
  real_num_collect:           "[| l is_const; m is_const |] ==>
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 ! *)
neuper@37974
   149
  real_num_collect_assoc:     "[| l is_const; m is_const |] ==> 
neuper@52148
   150
			      l * n + (m * n + k) = (l + m) * n + k" and
neuper@37974
   151
  real_num_collect_assoc_l:   "[| l is_const; m is_const |] ==>
neuper@37950
   152
			      l * n + (m * n + k) = (l + m)
neuper@52148
   153
				* n + k" and
neuper@37974
   154
  real_num_collect_assoc_r:   "[| l is_const; m is_const |] ==>
neuper@52148
   155
			      (k + m * n) + l * n = k + (l + m) * n" and
neuper@52148
   156
  real_one_collect:           "m is_const ==> 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 ! *)
neuper@52148
   159
  real_one_collect_assoc:     "m is_const ==> n + (m * n + k) = (1 + m)* n + k" and
neuper@37906
   160
neuper@52148
   161
  real_one_collect_assoc_l:   "m is_const ==> n + (m * n + k) = (1 + m) * n + k" and
neuper@52148
   162
  real_one_collect_assoc_r:  "m is_const ==> (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@60320
   172
  (*sym_real_mult_minus1 expands indefinitely without assumptions ...*)
walther@60320
   173
  real_mult_minus1_sym:   "[| \<not>(matches (- 1 * x) z); \<not>(z is_atom) |] ==> - (z::real) = -1 * z" and
wneuper@59587
   174
  real_mult_2:            "2 * z = z + (z::real)" and
wneuper@59587
   175
neuper@52148
   176
  real_add_mult_distrib_poly: "w is_polyexp ==> (z1 + z2) * w = z1 * w + z2 * w" and
neuper@37974
   177
  real_add_mult_distrib2_poly:"w is_polyexp ==> w * (z1 + z2) = w * z1 + w * z2"
neuper@37950
   178
walther@60278
   179
wneuper@59523
   180
subsection \<open>auxiliary functions\<close>
wneuper@59530
   181
ML \<open>
wneuper@59530
   182
val thy = @{theory};
wneuper@59530
   183
val poly_consts =
wneuper@59530
   184
  ["Groups.plus_class.plus", "Groups.minus_class.minus",
wneuper@59531
   185
  "Rings.divide_class.divide", "Groups.times_class.times",
walther@60275
   186
  "Transcendental.powr"];
wneuper@59530
   187
\<close>
wneuper@59523
   188
subsubsection \<open>for predicates in specifications (ML)\<close>
wneuper@59472
   189
ML \<open>
wneuper@59522
   190
(*--- auxiliary for is_expanded_in, is_poly_in, has_degree_in ---*)
walther@60317
   191
(*. a "monomial t in variable v" is a term t with
wneuper@59522
   192
  either (1) v NOT existent in t, or (2) v contained in t,
wneuper@59522
   193
  if (1) then degree 0
walther@60317
   194
  if (2) then v is a factor on the very right, casually with exponent.*)
wneuper@59522
   195
fun factor_right_deg (*case 2*)
walther@60317
   196
	    (Const ("Groups.times_class.times", _) $
walther@60317
   197
        t1 $ (Const ("Transcendental.powr",_) $ vv $ num)) v =
walther@60317
   198
	  if vv = v andalso not (Prog_Expr.occurs_in v t1) then SOME (snd (HOLogic.dest_number num))
walther@60317
   199
    else NONE
walther@60317
   200
  | factor_right_deg (Const ("Transcendental.powr",_) $ vv $ num) v =
walther@60317
   201
	   if (vv = v) then SOME (snd (HOLogic.dest_number num)) else NONE
wneuper@59522
   202
  | factor_right_deg (Const ("Groups.times_class.times",_) $ t1 $ vv) v = 
walther@59603
   203
	   if vv = v andalso not (Prog_Expr.occurs_in v t1) then SOME 1 else NONE
wneuper@59522
   204
  | factor_right_deg vv v =
wneuper@59522
   205
	  if (vv = v) then SOME 1 else NONE;    
wneuper@59522
   206
fun mono_deg_in m v =  (*case 1*)
walther@59603
   207
	if not (Prog_Expr.occurs_in v m) then (*case 1*) SOME 0 else factor_right_deg m v;
wneuper@59522
   208
wneuper@59522
   209
fun expand_deg_in t v =
wneuper@59522
   210
	let
wneuper@59522
   211
    fun edi ~1 ~1 (Const ("Groups.plus_class.plus", _) $ t1 $ t2) =
wneuper@59522
   212
          (case mono_deg_in t2 v of (* $ is left associative*)
wneuper@59522
   213
            SOME d' => edi d' d' t1 | NONE => NONE)
wneuper@59522
   214
      | edi ~1 ~1 (Const ("Groups.minus_class.minus", _) $ t1 $ t2) =
wneuper@59522
   215
          (case mono_deg_in t2 v of
wneuper@59522
   216
            SOME d' => edi d' d' t1 | NONE => NONE)
wneuper@59522
   217
      | edi d dmax (Const ("Groups.minus_class.minus", _) $ t1 $ t2) =
wneuper@59522
   218
          (case mono_deg_in t2 v of (*(d = 0 andalso d' = 0) handle 3+4-...4 +x*)
wneuper@59522
   219
	        SOME d' => if d > d' orelse (d = 0 andalso d' = 0) then edi d' dmax t1 else NONE
wneuper@59522
   220
          | NONE => NONE)
wneuper@59522
   221
      | edi d dmax (Const ("Groups.plus_class.plus",_) $ t1 $ t2) =
wneuper@59522
   222
          (case mono_deg_in t2 v of
wneuper@59522
   223
            SOME d' =>    (*RL (d = 0 andalso d' = 0) need to handle 3+4-...4 +x*)
wneuper@59522
   224
              if d > d' orelse (d = 0 andalso d' = 0) then edi d' dmax t1 else NONE
wneuper@59522
   225
          | NONE => NONE)
wneuper@59522
   226
      | edi ~1 ~1 t =
wneuper@59522
   227
          (case mono_deg_in t v of d as SOME _ => d | NONE => NONE)
wneuper@59522
   228
      | edi d dmax t = (*basecase last*)
wneuper@59522
   229
    	    (case mono_deg_in t v of
wneuper@59522
   230
    	      SOME d' => if d > d' orelse (d = 0 andalso d' = 0) then SOME dmax else NONE
wneuper@59522
   231
		      | NONE => NONE)
wneuper@59522
   232
	in edi ~1 ~1 t end;
wneuper@59522
   233
wneuper@59522
   234
fun poly_deg_in t v =
wneuper@59522
   235
	let
wneuper@59522
   236
    fun edi ~1 ~1 (Const ("Groups.plus_class.plus",_) $ t1 $ t2) =
wneuper@59522
   237
		    (case mono_deg_in t2 v of (* $ is left associative *)
wneuper@59522
   238
		      SOME d' => edi d' d' t1
wneuper@59522
   239
        | NONE => NONE)
wneuper@59522
   240
	    | edi d dmax (Const ("Groups.plus_class.plus",_) $ t1 $ t2) =
wneuper@59522
   241
		    (case mono_deg_in t2 v of
wneuper@59522
   242
	        SOME d' =>    (*RL (d = 0 andalso (d' = 0)) handle 3+4-...4 +x*)
wneuper@59522
   243
            if d > d' orelse (d = 0 andalso d' = 0) then edi d' dmax t1 else NONE
wneuper@59522
   244
        | NONE => NONE)
wneuper@59522
   245
	    | edi ~1 ~1 t =
wneuper@59522
   246
        (case mono_deg_in t v of
wneuper@59522
   247
		      d as SOME _ => d
wneuper@59522
   248
        | NONE => NONE)
wneuper@59522
   249
	    | edi d dmax t = (* basecase last *)
wneuper@59522
   250
		    (case mono_deg_in t v of
wneuper@59522
   251
		      SOME d' =>
wneuper@59522
   252
            if d > d' orelse (d = 0 andalso d' = 0) then SOME dmax else NONE
wneuper@59522
   253
        | NONE => NONE)
wneuper@59522
   254
	in edi ~1 ~1 t end;
wneuper@59523
   255
\<close>
neuper@37950
   256
wneuper@59523
   257
subsubsection \<open>for hard-coded AC rewriting (MG)\<close>
wneuper@59523
   258
ML \<open>
wneuper@59523
   259
(**. MG.03: make_polynomial_ ... uses SML-fun for ordering .**)
neuper@37950
   260
wneuper@59523
   261
(*FIXME.0401: make SML-order local to make_polynomial(_) *)
wneuper@59523
   262
(*FIXME.0401: replace 'make_polynomial'(old) by 'make_polynomial_'(MG) *)
wneuper@59523
   263
(* Polynom --> List von Monomen *) 
wneuper@59523
   264
fun poly2list (Const ("Groups.plus_class.plus",_) $ t1 $ t2) = 
wneuper@59523
   265
    (poly2list t1) @ (poly2list t2)
wneuper@59523
   266
  | poly2list t = [t];
neuper@37950
   267
wneuper@59523
   268
(* Monom --> Liste von Variablen *)
wneuper@59523
   269
fun monom2list (Const ("Groups.times_class.times",_) $ t1 $ t2) = 
wneuper@59523
   270
    (monom2list t1) @ (monom2list t2)
wneuper@59523
   271
  | monom2list t = [t];
neuper@37950
   272
walther@60318
   273
(* use this fun ONLY locally: makes negative numbers with "-" instead of "~" for poly-orders*)
walther@60318
   274
fun dest_number' t =
walther@60318
   275
  let
walther@60318
   276
    val i = TermC.num_of_term t
walther@60318
   277
  in
walther@60318
   278
    if i >= 0 then HOLogic.dest_number t |> snd |> string_of_int
walther@60318
   279
    else (i * ~1) |> string_of_int |> curry op ^ "-"
walther@60318
   280
  end
walther@60318
   281
wneuper@59523
   282
(* liefert Variablenname (String) einer Variablen und Basis bei Potenz *)
walther@60275
   283
fun get_basStr (Const ("Transcendental.powr",_) $ Free (str, _) $ _) = str
walther@60318
   284
  | get_basStr (Const ("Transcendental.powr",_) $ n $ _) = dest_number' n
wneuper@59523
   285
  | get_basStr (Free (str, _)) = str
walther@60318
   286
  | get_basStr t =
walther@60318
   287
    if TermC.is_num t then dest_number' t
walther@60318
   288
    else "|||"; (* gross gewichtet; für Brüche ect. *)
neuper@37950
   289
wneuper@59523
   290
(* liefert Hochzahl (String) einer Variablen bzw Gewichtstring (zum Sortieren) *)
walther@60318
   291
fun get_potStr (Const ("Transcendental.powr",_) $ Free _ $ Free (str, _)) = str
walther@60318
   292
  | get_potStr (Const ("Transcendental.powr",_) $ Free _ $ t) =
walther@60318
   293
    if TermC.is_num t then dest_number' t else "|||"
wneuper@59523
   294
  | get_potStr (Free (_, _)) = "---" (* keine Hochzahl --> kleinst gewichtet *)
wneuper@59523
   295
  | get_potStr _ = "||||||"; (* gross gewichtet; für Brüch ect. *)
neuper@37978
   296
wneuper@59523
   297
(* Umgekehrte string_ord *)
wneuper@59523
   298
val string_ord_rev =  rev_order o string_ord;
wneuper@59523
   299
		
wneuper@59523
   300
 (* Ordnung zum lexikographischen Vergleich zweier Variablen (oder Potenzen) 
wneuper@59523
   301
    innerhalb eines Monomes:
wneuper@59523
   302
    - zuerst lexikographisch nach Variablenname 
wneuper@59523
   303
    - wenn gleich: nach steigender Potenz *)
walther@60318
   304
fun var_ord (a,b: term) = 
walther@60318
   305
  prod_ord string_ord string_ord 
wneuper@59523
   306
    ((get_basStr a, get_potStr a), (get_basStr b, get_potStr b));
neuper@37950
   307
wneuper@59523
   308
(* Ordnung zum lexikographischen Vergleich zweier Variablen (oder Potenzen); 
wneuper@59523
   309
   verwendet zum Sortieren von Monomen mittels Gesamtgradordnung:
wneuper@59523
   310
   - zuerst lexikographisch nach Variablenname 
wneuper@59523
   311
   - wenn gleich: nach sinkender Potenz*)
walther@60318
   312
fun var_ord_revPow (a, b: term) =
walther@60318
   313
  prod_ord string_ord string_ord_rev 
wneuper@59523
   314
    ((get_basStr a, get_potStr a), (get_basStr b, get_potStr b));
neuper@37950
   315
wneuper@59523
   316
(* Ordnet ein Liste von Variablen (und Potenzen) lexikographisch *)
wneuper@59523
   317
val sort_varList = sort var_ord;
wneuper@59523
   318
wneuper@59523
   319
(* Entfernet aeussersten Operator (Wurzel) aus einem Term und schreibt 
wneuper@59523
   320
   Argumente in eine Liste *)
wneuper@59523
   321
fun args u : term list =
walther@60318
   322
  let
walther@60318
   323
    fun stripc (f $ t, ts) = stripc (f, t::ts)
walther@60318
   324
  	  | stripc (t as Free _, ts) = (t::ts)
walther@60318
   325
  	  | stripc (_, ts) = ts
walther@60318
   326
  in stripc (u, []) end;
wneuper@59523
   327
                                    
wneuper@59523
   328
(* liefert True, falls der Term (Liste von Termen) nur Zahlen 
wneuper@59523
   329
   (keine Variablen) enthaelt *)
walther@60317
   330
fun filter_num ts = fold (curry and_) (map TermC.is_num ts) true
wneuper@59523
   331
wneuper@59523
   332
(* liefert True, falls der Term nur Zahlen (keine Variablen) enthaelt 
wneuper@59523
   333
   dh. er ist ein numerischer Wert und entspricht einem Koeffizienten *)
wneuper@59523
   334
fun is_nums t = filter_num [t];
wneuper@59523
   335
wneuper@59523
   336
(* Berechnet den Gesamtgrad eines Monoms *)
walther@60318
   337
(**)local(**)
walther@60317
   338
  fun counter (n, []) = n
walther@60317
   339
    | counter (n, x :: xs) = 
walther@60317
   340
	    if (is_nums x) then counter (n, xs)
walther@60317
   341
	    else 
walther@60317
   342
	      (case x of 
walther@60318
   343
		      (Const ("Transcendental.powr", _) $ Free _ $ t) =>
walther@60318
   344
            if TermC.is_num t
walther@60318
   345
            then counter (t |> HOLogic.dest_number |> snd |> curry op + n, xs)
walther@60318
   346
            else counter (n + 1000, xs) (*FIXME.MG?!*)
walther@60318
   347
	      | (Const ("Num.numeral_class.numeral", _) $ num) =>
walther@60318
   348
            counter (n + 1 + HOLogic.dest_numeral num, xs)
walther@60318
   349
	      | _ => counter (n + 1, xs)) (*FIXME.MG?! ... Brüche ect.*)
walther@60318
   350
(**)in(**)
walther@60317
   351
  fun monom_degree l = counter (0, l) 
walther@60318
   352
(**)end;(*local*)
wneuper@59523
   353
wneuper@59523
   354
(* wie Ordnung dict_ord (lexicographische Ordnung zweier Listen, mit Vergleich 
wneuper@59523
   355
   der Listen-Elemente mit elem_ord) - Elemente die Bedingung cond erfuellen, 
wneuper@59523
   356
   werden jedoch dabei ignoriert (uebersprungen)  *)
wneuper@59523
   357
fun dict_cond_ord _ _ ([], []) = EQUAL
wneuper@59523
   358
  | dict_cond_ord _ _ ([], _ :: _) = LESS
wneuper@59523
   359
  | dict_cond_ord _ _ (_ :: _, []) = GREATER
wneuper@59523
   360
  | dict_cond_ord elem_ord cond (x :: xs, y :: ys) =
wneuper@59523
   361
    (case (cond x, cond y) of 
walther@60318
   362
	    (false, false) =>
walther@60318
   363
        (case elem_ord (x, y) of 
walther@60318
   364
				  EQUAL => dict_cond_ord elem_ord cond (xs, ys) 
walther@60318
   365
			  | ord => ord)
walther@60318
   366
    | (false, true)  => dict_cond_ord elem_ord cond (x :: xs, ys)
walther@60318
   367
    | (true, false)  => dict_cond_ord elem_ord cond (xs, y :: ys)
walther@60318
   368
    | (true, true)  =>  dict_cond_ord elem_ord cond (xs, ys) );
wneuper@59523
   369
wneuper@59523
   370
(* Gesamtgradordnung zum Vergleich von Monomen (Liste von Variablen/Potenzen):
wneuper@59523
   371
   zuerst nach Gesamtgrad, bei gleichem Gesamtgrad lexikographisch ordnen - 
walther@60260
   372
   dabei werden Koeffizienten ignoriert (2*3*a \<up> 2*4*b gilt wie a \<up> 2*b) *)
wneuper@59523
   373
fun degree_ord (xs, ys) =
walther@60318
   374
	prod_ord int_ord (dict_cond_ord var_ord_revPow is_nums)
walther@60318
   375
	  ((monom_degree xs, xs), (monom_degree ys, ys));
wneuper@59523
   376
wneuper@59523
   377
fun hd_str str = substring (str, 0, 1);
wneuper@59523
   378
fun tl_str str = substring (str, 1, (size str) - 1);
wneuper@59523
   379
wneuper@59523
   380
(* liefert nummerischen Koeffizienten eines Monoms oder NONE *)
walther@60318
   381
fun get_koeff_of_mon [] = raise ERROR "get_koeff_of_mon: called with l = []"
walther@60318
   382
  | get_koeff_of_mon (x :: _) = if is_nums x then SOME x else NONE;
wneuper@59523
   383
wneuper@59523
   384
(* wandelt Koeffizient in (zum sortieren geeigneten) String um *)
walther@60318
   385
fun koeff2ordStr (SOME t) =
walther@60318
   386
    if TermC.is_num t
walther@60318
   387
    then 
walther@60318
   388
      if (t |> HOLogic.dest_number |> snd) < 0
walther@60318
   389
      then (t |> HOLogic.dest_number |> snd |> curry op * ~1 |> string_of_int) ^ "0"  (* 3 < -3 *)
walther@60318
   390
      else (t |> HOLogic.dest_number |> snd |> string_of_int)
walther@60318
   391
    else "aaa"                                                      (* "num.Ausdruck" --> gross *)
walther@60318
   392
  | koeff2ordStr NONE = "---";                                     (* "kein Koeff" --> kleinste *)
wneuper@59523
   393
wneuper@59523
   394
(* Order zum Vergleich von Koeffizienten (strings): 
wneuper@59523
   395
   "kein Koeff" < "0" < "1" < "-1" < "2" < "-2" < ... < "num.Ausdruck" *)
walther@60318
   396
fun compare_koeff_ord (xs, ys) = string_ord
walther@60318
   397
  ((koeff2ordStr o get_koeff_of_mon) xs,
walther@60318
   398
   (koeff2ordStr o get_koeff_of_mon) ys);
wneuper@59523
   399
wneuper@59523
   400
(* Gesamtgradordnung degree_ord + Ordnen nach Koeffizienten falls EQUAL *)
wneuper@59523
   401
fun koeff_degree_ord (xs, ys) =
wneuper@59523
   402
	    prod_ord degree_ord compare_koeff_ord ((xs, xs), (ys, ys));
wneuper@59523
   403
wneuper@59523
   404
(* Ordnet ein Liste von Monomen (Monom = Liste von Variablen) mittels 
wneuper@59523
   405
   Gesamtgradordnung *)
wneuper@59523
   406
val sort_monList = sort koeff_degree_ord;
wneuper@59523
   407
wneuper@59523
   408
(* Alternativ zu degree_ord koennte auch die viel einfachere und 
wneuper@59523
   409
   kuerzere Ordnung simple_ord verwendet werden - ist aber nicht 
wneuper@59523
   410
   fuer unsere Zwecke geeignet!
wneuper@59523
   411
wneuper@59523
   412
fun simple_ord (al,bl: term list) = dict_ord string_ord 
wneuper@59523
   413
	 (map get_basStr al, map get_basStr bl); 
wneuper@59523
   414
wneuper@59523
   415
val sort_monList = sort simple_ord; *)
wneuper@59523
   416
wneuper@59523
   417
(* aus 2 Variablen wird eine Summe bzw ein Produkt erzeugt 
wneuper@59523
   418
   (mit gewuenschtem Typen T) *)
wneuper@59523
   419
fun plus T = Const ("Groups.plus_class.plus", [T,T] ---> T);
wneuper@59523
   420
fun mult T = Const ("Groups.times_class.times", [T,T] ---> T);
wneuper@59523
   421
fun binop op_ t1 t2 = op_ $ t1 $ t2;
wneuper@59523
   422
fun create_prod T (a,b) = binop (mult T) a b;
wneuper@59523
   423
fun create_sum T (a,b) = binop (plus T) a b;
wneuper@59523
   424
wneuper@59523
   425
(* löscht letztes Element einer Liste *)
wneuper@59523
   426
fun drop_last l = take ((length l)-1,l);
wneuper@59523
   427
wneuper@59523
   428
(* Liste von Variablen --> Monom *)
wneuper@59523
   429
fun create_monom T vl = foldr (create_prod T) (drop_last vl, last_elem vl);
wneuper@59523
   430
(* Bemerkung: 
wneuper@59523
   431
   foldr bewirkt rechtslastige Klammerung des Monoms - ist notwendig, damit zwei 
wneuper@59523
   432
   gleiche Monome zusammengefasst werden können (collect_numerals)! 
wneuper@59523
   433
   zB: 2*(x*(y*z)) + 3*(x*(y*z)) --> (2+3)*(x*(y*z))*)
wneuper@59523
   434
wneuper@59523
   435
(* Liste von Monomen --> Polynom *)	
wneuper@59523
   436
fun create_polynom T ml = foldl (create_sum T) (hd ml, tl ml);
wneuper@59523
   437
(* Bemerkung: 
wneuper@59523
   438
   foldl bewirkt linkslastige Klammerung des Polynoms (der Summanten) - 
wneuper@59523
   439
   bessere Darstellung, da keine Klammern sichtbar! 
wneuper@59523
   440
   (und discard_parentheses in make_polynomial hat weniger zu tun) *)
wneuper@59523
   441
wneuper@59523
   442
(* sorts the variables (faktors) of an expanded polynomial lexicographical *)
wneuper@59523
   443
fun sort_variables t = 
walther@60317
   444
  let
walther@60317
   445
  	val ll = map monom2list (poly2list t);
walther@60317
   446
  	val lls = map sort_varList ll; 
walther@60317
   447
  	val T = type_of t;
walther@60317
   448
  	val ls = map (create_monom T) lls;
walther@60317
   449
  in create_polynom T ls end;
wneuper@59523
   450
wneuper@59523
   451
(* sorts the monoms of an expanded and variable-sorted polynomial 
wneuper@59523
   452
   by total_degree *)
wneuper@59523
   453
fun sort_monoms t = 
walther@60318
   454
  let
walther@60318
   455
  	val ll =  map monom2list (poly2list t);
walther@60318
   456
  	val lls = sort_monList ll;
walther@60318
   457
  	val T = Term.type_of t;
walther@60318
   458
  	val ls = map (create_monom T) lls;
walther@60318
   459
  in create_polynom T ls end;
wneuper@59523
   460
\<close>
wneuper@59523
   461
wneuper@59523
   462
subsubsection \<open>rewrite order for hard-coded AC rewriting\<close>
wneuper@59523
   463
ML \<open>
neuper@37950
   464
local (*. for make_polynomial .*)
neuper@37950
   465
neuper@37950
   466
open Term;  (* for type order = EQUAL | LESS | GREATER *)
neuper@37950
   467
neuper@37950
   468
fun pr_ord EQUAL = "EQUAL"
neuper@37950
   469
  | pr_ord LESS  = "LESS"
neuper@37950
   470
  | pr_ord GREATER = "GREATER";
neuper@37950
   471
neuper@37950
   472
fun dest_hd' (Const (a, T)) =                          (* ~ term.ML *)
neuper@37950
   473
  (case a of
walther@60275
   474
     "Transcendental.powr" => ((("|||||||||||||", 0), T), 0)    (*WN greatest string*)
neuper@37950
   475
   | _ => (((a, 0), T), 0))
walther@60317
   476
  | dest_hd' (Free (a, T)) = (((a, 0), T), 1)(*TODOO handle this as numeral, too? see EqSystem.thy*)
neuper@37950
   477
  | dest_hd' (Var v) = (v, 2)
neuper@37950
   478
  | dest_hd' (Bound i) = ((("", i), dummyT), 3)
wneuper@59523
   479
  | dest_hd' (Abs (_, T, _)) = ((("", 0), T), 4)
wneuper@59523
   480
  | dest_hd' t = raise TERM ("dest_hd'", [t]);
neuper@37950
   481
neuper@37950
   482
fun size_of_term' (Const(str,_) $ t) =
walther@60275
   483
  if "Transcendental.powr"= str then 1000 + size_of_term' t else 1+size_of_term' t(*WN*)
neuper@37950
   484
  | size_of_term' (Abs (_,_,body)) = 1 + size_of_term' body
neuper@37950
   485
  | size_of_term' (f$t) = size_of_term' f  +  size_of_term' t
neuper@37950
   486
  | size_of_term' _ = 1;
neuper@37950
   487
neuper@37950
   488
fun term_ord' pr thy (Abs (_, T, t), Abs(_, U, u)) =       (* ~ term.ML *)
neuper@52070
   489
    (case term_ord' pr thy (t, u) of EQUAL => Term_Ord.typ_ord (T, U) | ord => ord)
neuper@37950
   490
  | term_ord' pr thy (t, u) =
neuper@52070
   491
    (if pr then 
neuper@52070
   492
	   let
neuper@52070
   493
       val (f, ts) = strip_comb t and (g, us) = strip_comb u;
walther@59870
   494
       val _ = tracing ("t= f@ts= \"" ^ UnparseC.term_in_thy thy f ^ "\" @ \"[" ^
walther@59870
   495
         commas (map (UnparseC.term_in_thy thy) ts) ^ "]\"");
walther@59870
   496
       val _ = tracing("u= g@us= \"" ^ UnparseC.term_in_thy thy g ^ "\" @ \"[" ^
walther@59870
   497
         commas (map (UnparseC.term_in_thy thy) us) ^ "]\"");
neuper@52070
   498
       val _ = tracing ("size_of_term(t,u)= (" ^ string_of_int (size_of_term' t) ^ ", " ^
neuper@52070
   499
         string_of_int (size_of_term' u) ^ ")");
neuper@52070
   500
       val _ = tracing ("hd_ord(f,g)      = " ^ (pr_ord o hd_ord) (f,g));
neuper@52070
   501
       val _ = tracing ("terms_ord(ts,us) = " ^ (pr_ord o terms_ord str false) (ts, us));
neuper@52070
   502
       val _ = tracing ("-------");
neuper@52070
   503
     in () end
neuper@37950
   504
       else ();
neuper@37950
   505
	 case int_ord (size_of_term' t, size_of_term' u) of
neuper@37950
   506
	   EQUAL =>
neuper@37950
   507
	     let val (f, ts) = strip_comb t and (g, us) = strip_comb u in
neuper@37950
   508
	       (case hd_ord (f, g) of EQUAL => (terms_ord str pr) (ts, us) 
neuper@37950
   509
	     | ord => ord)
neuper@37950
   510
	     end
neuper@37950
   511
	 | ord => ord)
neuper@37950
   512
and hd_ord (f, g) =                                        (* ~ term.ML *)
neuper@37974
   513
  prod_ord (prod_ord Term_Ord.indexname_ord Term_Ord.typ_ord) int_ord (dest_hd' f, dest_hd' g)
wneuper@59523
   514
and terms_ord _ pr (ts, us) = 
walther@59881
   515
    list_ord (term_ord' pr (ThyC.get_theory "Isac_Knowledge"))(ts, us);
neuper@52070
   516
neuper@37950
   517
in
neuper@37950
   518
walther@59910
   519
fun ord_make_polynomial (pr:bool) thy (_: subst) tu = 
neuper@37950
   520
    (term_ord' pr thy(***) tu = LESS );
neuper@37950
   521
neuper@37950
   522
end;(*local*)
neuper@37950
   523
walther@59857
   524
Rewrite_Ord.rew_ord' := overwritel (! Rewrite_Ord.rew_ord', (* TODO: make analogous to KEStore_Elems.add_mets *)
wneuper@59523
   525
[("termlessI", termlessI), ("ord_make_polynomial", ord_make_polynomial false thy)]);
wneuper@59523
   526
\<close>
neuper@37950
   527
wneuper@59523
   528
subsection \<open>predicates\<close>
wneuper@59523
   529
subsubsection \<open>in specifications\<close>
wneuper@59523
   530
ML \<open>
wneuper@59523
   531
(* is_polyrat_in becomes true, if no bdv is in the denominator of a fraction*)
wneuper@59523
   532
fun is_polyrat_in t v = 
wneuper@59524
   533
  let
walther@59962
   534
   	fun finddivide (_ $ _ $ _ $ _) _ = raise ERROR("is_polyrat_in:")
wneuper@59523
   535
	    (* at the moment there is no term like this, but ....*)
walther@59603
   536
	  | finddivide (Const ("Rings.divide_class.divide",_) $ _ $ b) v = not (Prog_Expr.occurs_in v b)
wneuper@59524
   537
	  | finddivide (_ $ t1 $ t2) v = finddivide t1 v orelse finddivide t2 v
wneuper@59524
   538
	  | finddivide (_ $ t1) v = finddivide t1 v
wneuper@59523
   539
	  | finddivide _ _ = false;
wneuper@59524
   540
  in finddivide t v end;
wneuper@59523
   541
    
wneuper@59524
   542
fun is_expanded_in t v = case expand_deg_in t v of SOME _ => true | NONE => false;
wneuper@59524
   543
fun is_poly_in t v =     case poly_deg_in t v of SOME _ => true | NONE => false;
wneuper@59524
   544
fun has_degree_in t v =  case expand_deg_in t v of SOME d => d | NONE => ~1;
neuper@37950
   545
wneuper@59523
   546
(*.the expression contains + - * ^ only ?
wneuper@59523
   547
   this is weaker than 'is_polynomial' !.*)
wneuper@59523
   548
fun is_polyexp (Free _) = true
wneuper@59529
   549
  | is_polyexp (Const _) = true (* potential danger: bdv is not considered *)
walther@60318
   550
  | is_polyexp (Const ("Groups.plus_class.plus",_) $ Free _ $ num) =
walther@60318
   551
    if TermC.is_num num then true
walther@60318
   552
    else if TermC.is_variable num then true
walther@60318
   553
    else is_polyexp num
walther@60318
   554
  | is_polyexp (Const ("Groups.plus_class.plus",_) $ num $ Free _) =
walther@60318
   555
    if TermC.is_num num then true
walther@60318
   556
    else if TermC.is_variable num then true
walther@60318
   557
    else is_polyexp num
walther@60318
   558
  | is_polyexp (Const ("Groups.minus_class.minus",_) $ Free _ $ num) =
walther@60318
   559
    if TermC.is_num num then true
walther@60318
   560
    else if TermC.is_variable num then true
walther@60318
   561
    else is_polyexp num
walther@60318
   562
  | is_polyexp (Const ("Groups.times_class.times",_) $ num $ Free _) =
walther@60318
   563
    if TermC.is_num num then true
walther@60318
   564
    else if TermC.is_variable num then true
walther@60318
   565
    else is_polyexp num
walther@60318
   566
  | is_polyexp (Const ("Transcendental.powr",_) $ Free _ $ num) =
walther@60318
   567
    if TermC.is_num num then true
walther@60318
   568
    else if TermC.is_variable num then true
walther@60318
   569
    else is_polyexp num
wneuper@59523
   570
  | is_polyexp (Const ("Groups.plus_class.plus",_) $ t1 $ t2) = 
walther@60318
   571
    ((is_polyexp t1) andalso (is_polyexp t2))
wneuper@59523
   572
  | is_polyexp (Const ("Groups.minus_class.minus",_) $ t1 $ t2) = 
walther@60318
   573
    ((is_polyexp t1) andalso (is_polyexp t2))
wneuper@59523
   574
  | is_polyexp (Const ("Groups.times_class.times",_) $ t1 $ t2) = 
walther@60318
   575
    ((is_polyexp t1) andalso (is_polyexp t2))
walther@60275
   576
  | is_polyexp (Const ("Transcendental.powr",_) $ t1 $ t2) = 
walther@60318
   577
    ((is_polyexp t1) andalso (is_polyexp t2))
walther@60317
   578
  | is_polyexp num = TermC.is_num num;
wneuper@59523
   579
\<close>
neuper@37950
   580
wneuper@59523
   581
subsubsection \<open>for hard-coded AC rewriting\<close>
wneuper@59523
   582
ML \<open>
wneuper@59523
   583
(* auch Klammerung muss übereinstimmen;
wneuper@59523
   584
   sort_variables klammert Produkte rechtslastig*)
wneuper@59523
   585
fun is_multUnordered t = ((is_polyexp t) andalso not (t = sort_variables t));
wneuper@59523
   586
wneuper@59523
   587
fun is_addUnordered t = ((is_polyexp t) andalso not (t = sort_monoms t));
wneuper@59523
   588
\<close>
wneuper@59523
   589
wneuper@59523
   590
subsection \<open>evaluations functions\<close>
walther@60318
   591
subsubsection \<open>for predicates\<close>                   
wneuper@59523
   592
ML \<open>
walther@60278
   593
fun eval_is_polyrat_in _ _(p as (Const ("Poly.is_polyrat_in",_) $ t $ v)) _  =
wneuper@59523
   594
    if is_polyrat_in t v 
walther@59868
   595
    then SOME ((UnparseC.term p) ^ " = True",
wneuper@59523
   596
	        HOLogic.Trueprop $ (TermC.mk_equality (p, @{term True})))
walther@59868
   597
    else SOME ((UnparseC.term p) ^ " = True",
wneuper@59523
   598
	        HOLogic.Trueprop $ (TermC.mk_equality (p, @{term False})))
wneuper@59523
   599
  | eval_is_polyrat_in _ _ _ _ = ((*tracing"### no matches";*) NONE);
wneuper@59523
   600
walther@60278
   601
(*("is_expanded_in", ("Poly.is_expanded_in", eval_is_expanded_in ""))*)
wneuper@59523
   602
fun eval_is_expanded_in _ _ 
walther@60278
   603
       (p as (Const ("Poly.is_expanded_in",_) $ t $ v)) _ =
wneuper@59523
   604
    if is_expanded_in t v
walther@59868
   605
    then SOME ((UnparseC.term p) ^ " = True",
wneuper@59523
   606
	        HOLogic.Trueprop $ (TermC.mk_equality (p, @{term True})))
walther@59868
   607
    else SOME ((UnparseC.term p) ^ " = True",
wneuper@59523
   608
	        HOLogic.Trueprop $ (TermC.mk_equality (p, @{term False})))
wneuper@59523
   609
  | eval_is_expanded_in _ _ _ _ = NONE;
wneuper@59523
   610
walther@60278
   611
(*("is_poly_in", ("Poly.is_poly_in", eval_is_poly_in ""))*)
wneuper@59523
   612
fun eval_is_poly_in _ _ 
walther@60278
   613
       (p as (Const ("Poly.is_poly_in",_) $ t $ v)) _ =
wneuper@59523
   614
    if is_poly_in t v
walther@59868
   615
    then SOME ((UnparseC.term p) ^ " = True",
wneuper@59523
   616
	        HOLogic.Trueprop $ (TermC.mk_equality (p, @{term True})))
walther@59868
   617
    else SOME ((UnparseC.term p) ^ " = True",
wneuper@59523
   618
	        HOLogic.Trueprop $ (TermC.mk_equality (p, @{term False})))
wneuper@59523
   619
  | eval_is_poly_in _ _ _ _ = NONE;
wneuper@59523
   620
walther@60278
   621
(*("has_degree_in", ("Poly.has_degree_in", eval_has_degree_in ""))*)
wneuper@59523
   622
fun eval_has_degree_in _ _ 
walther@60278
   623
	     (p as (Const ("Poly.has_degree_in",_) $ t $ v)) _ =
wneuper@59523
   624
    let val d = has_degree_in t v
wneuper@59523
   625
	val d' = TermC.term_of_num HOLogic.realT d
walther@59868
   626
    in SOME ((UnparseC.term p) ^ " = " ^ (string_of_int d),
wneuper@59523
   627
	      HOLogic.Trueprop $ (TermC.mk_equality (p, d')))
wneuper@59523
   628
    end
wneuper@59523
   629
  | eval_has_degree_in _ _ _ _ = NONE;
wneuper@59523
   630
walther@60278
   631
(*("is_polyexp", ("Poly.is_polyexp", eval_is_polyexp ""))*)
wneuper@59523
   632
fun eval_is_polyexp (thmid:string) _ 
walther@60278
   633
		       (t as (Const("Poly.is_polyexp", _) $ arg)) thy = 
wneuper@59523
   634
    if is_polyexp arg
walther@59870
   635
    then SOME (TermC.mk_thmid thmid (UnparseC.term_in_thy thy arg) "", 
wneuper@59523
   636
	         HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True})))
walther@59870
   637
    else SOME (TermC.mk_thmid thmid (UnparseC.term_in_thy thy arg) "", 
wneuper@59523
   638
	         HOLogic.Trueprop $ (TermC.mk_equality (t, @{term False})))
wneuper@59523
   639
  | eval_is_polyexp _ _ _ _ = NONE; 
wneuper@59523
   640
\<close>
wneuper@59523
   641
wneuper@59523
   642
subsubsection \<open>for hard-coded AC rewriting\<close>
wneuper@59523
   643
ML \<open>
wneuper@59523
   644
(*WN.18.6.03 *)
walther@60278
   645
(*("is_addUnordered", ("Poly.is_addUnordered", eval_is_addUnordered ""))*)
wneuper@59523
   646
fun eval_is_addUnordered (thmid:string) _ 
walther@60278
   647
		       (t as (Const("Poly.is_addUnordered", _) $ arg)) thy = 
wneuper@59523
   648
    if is_addUnordered arg
walther@59870
   649
    then SOME (TermC.mk_thmid thmid (UnparseC.term_in_thy thy arg) "", 
wneuper@59523
   650
	         HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True})))
walther@59870
   651
    else SOME (TermC.mk_thmid thmid (UnparseC.term_in_thy thy arg) "", 
wneuper@59523
   652
	         HOLogic.Trueprop $ (TermC.mk_equality (t, @{term False})))
wneuper@59523
   653
  | eval_is_addUnordered _ _ _ _ = NONE; 
wneuper@59523
   654
wneuper@59523
   655
fun eval_is_multUnordered (thmid:string) _ 
walther@60278
   656
		       (t as (Const("Poly.is_multUnordered", _) $ arg)) thy = 
wneuper@59523
   657
    if is_multUnordered arg
walther@59870
   658
    then SOME (TermC.mk_thmid thmid (UnparseC.term_in_thy thy arg) "", 
wneuper@59523
   659
	         HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True})))
walther@59870
   660
    else SOME (TermC.mk_thmid thmid (UnparseC.term_in_thy thy arg) "", 
wneuper@59523
   661
	         HOLogic.Trueprop $ (TermC.mk_equality (t, @{term False})))
wneuper@59523
   662
  | eval_is_multUnordered _ _ _ _ = NONE; 
wneuper@59523
   663
\<close>
wneuper@59526
   664
setup \<open>KEStore_Elems.add_calcs
walther@60278
   665
  [("is_polyrat_in", ("Poly.is_polyrat_in",
wneuper@59526
   666
		    eval_is_polyrat_in "#eval_is_polyrat_in")),
walther@60278
   667
    ("is_expanded_in", ("Poly.is_expanded_in", eval_is_expanded_in "")),
walther@60278
   668
    ("is_poly_in", ("Poly.is_poly_in", eval_is_poly_in "")),
walther@60278
   669
    ("has_degree_in", ("Poly.has_degree_in", eval_has_degree_in "")),
walther@60278
   670
    ("is_polyexp", ("Poly.is_polyexp", eval_is_polyexp "")),
walther@60278
   671
    ("is_multUnordered", ("Poly.is_multUnordered", eval_is_multUnordered"")),
walther@60278
   672
    ("is_addUnordered", ("Poly.is_addUnordered", eval_is_addUnordered ""))]\<close>
wneuper@59523
   673
wneuper@59523
   674
subsection \<open>rule-sets\<close>
wneuper@59523
   675
subsubsection \<open>without specific order\<close>
wneuper@59523
   676
ML \<open>
wneuper@59523
   677
(* used only for merge *)
walther@59852
   678
val calculate_Poly = Rule_Set.append_rules "calculate_PolyFIXXXME.not.impl." Rule_Set.empty [];
wneuper@59523
   679
wneuper@59523
   680
(*.for evaluation of conditions in rewrite rules.*)
walther@59852
   681
val Poly_erls = Rule_Set.append_rules "Poly_erls" Atools_erls
walther@59878
   682
  [Rule.Eval ("HOL.eq", Prog_Expr.eval_equal "#equal_"),
walther@59871
   683
  Rule.Thm  ("real_unari_minus", ThmC.numerals_to_Free @{thm real_unari_minus}),
walther@60278
   684
  Rule.Eval ("Groups.plus_class.plus", eval_binop "#add_"),
walther@60278
   685
  Rule.Eval ("Groups.minus_class.minus", eval_binop "#sub_"),
walther@60278
   686
  Rule.Eval ("Groups.times_class.times", eval_binop "#mult_"),
walther@60278
   687
  Rule.Eval ("Transcendental.powr", eval_binop "#power_")];
wneuper@59523
   688
walther@59852
   689
val poly_crls = Rule_Set.append_rules "poly_crls" Atools_crls
walther@59878
   690
  [Rule.Eval ("HOL.eq", Prog_Expr.eval_equal "#equal_"),
walther@59871
   691
  Rule.Thm ("real_unari_minus", ThmC.numerals_to_Free @{thm real_unari_minus}),
walther@60278
   692
  Rule.Eval ("Groups.plus_class.plus", eval_binop "#add_"),
walther@60278
   693
  Rule.Eval ("Groups.minus_class.minus", eval_binop "#sub_"),
walther@60278
   694
  Rule.Eval ("Groups.times_class.times", eval_binop "#mult_"),
walther@60278
   695
  Rule.Eval ("Transcendental.powr" , eval_binop "#power_")];
wneuper@59523
   696
\<close>
wneuper@59523
   697
ML \<open>
neuper@37950
   698
val expand =
walther@59857
   699
  Rule_Def.Repeat {id = "expand", preconds = [], rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
walther@59852
   700
      erls = Rule_Set.empty,srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@59871
   701
      rules = [Rule.Thm ("distrib_right" , ThmC.numerals_to_Free @{thm distrib_right}),
neuper@37950
   702
	       (*"(z1.0 + z2.0) * w = z1.0 * w + z2.0 * w"*)
walther@59871
   703
	       Rule.Thm ("distrib_left", ThmC.numerals_to_Free @{thm distrib_left})
neuper@37950
   704
	       (*"w * (z1.0 + z2.0) = w * z1.0 + w * z2.0"*)
walther@59878
   705
	       ], scr = Rule.Empty_Prog};
neuper@37950
   706
walther@60320
   707
(* erls for calculate_Rational + etc *)
walther@60318
   708
val powers_erls =
walther@60318
   709
  Rule_Def.Repeat {id = "powers_erls", preconds = [], rew_ord = ("dummy_ord",Rewrite_Ord.dummy_ord), 
walther@60318
   710
      erls = Rule_Set.empty, srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@60320
   711
      rules = 
walther@60320
   712
        [Rule.Eval ("Prog_Expr.matches", Prog_Expr.eval_matches "#matches_"),
walther@60320
   713
	       Rule.Eval ("Prog_Expr.is_atom", Prog_Expr.eval_is_atom "#is_atom_"),
walther@60318
   714
	       Rule.Eval ("Prog_Expr.is_even", Prog_Expr.eval_is_even "#is_even_"),
walther@60318
   715
	       Rule.Eval ("Orderings.ord_class.less", Prog_Expr.eval_equ "#less_"),
walther@60318
   716
	       Rule.Thm ("not_false", ThmC.numerals_to_Free @{thm not_false}),
walther@60318
   717
	       Rule.Thm ("not_true", ThmC.numerals_to_Free @{thm not_true}),
walther@60318
   718
	       Rule.Eval ("Groups.plus_class.plus", (**)eval_binop "#add_")
walther@60318
   719
	       ],
walther@60318
   720
      scr = Rule.Empty_Prog
walther@60318
   721
      };
walther@60320
   722
walther@60320
   723
val discard_minus =
walther@60320
   724
  Rule_Def.Repeat {id = "discard_minus", preconds = [], rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
walther@60320
   725
      erls = powers_erls, srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@60320
   726
      rules =
walther@60320
   727
       [Rule.Thm ("real_diff_minus", ThmC.numerals_to_Free @{thm real_diff_minus}),
walther@60320
   728
          (*"a - b = a + -1 * b"*)
walther@60320
   729
        Rule.Thm ("real_mult_minus1_sym", ThmC.numerals_to_Free (@{thm real_mult_minus1_sym}))
walther@60320
   730
	        (*"\<not>(z is_const) ==> - (z::real) = -1 * z"*)],
walther@60320
   731
	      scr = Rule.Empty_Prog};
walther@60320
   732
neuper@37950
   733
val expand_poly_ = 
walther@59851
   734
  Rule_Def.Repeat{id = "expand_poly_", preconds = [], 
walther@59857
   735
      rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
walther@60318
   736
      erls = powers_erls, srls = Rule_Set.Empty,
neuper@42451
   737
      calc = [], errpatts = [],
neuper@42407
   738
      rules =
walther@59871
   739
        [Rule.Thm ("real_plus_binom_pow4", ThmC.numerals_to_Free @{thm real_plus_binom_pow4}),
walther@60260
   740
	           (*"(a + b) \<up> 4 = ... "*)
walther@59871
   741
	         Rule.Thm ("real_plus_binom_pow5",ThmC.numerals_to_Free @{thm real_plus_binom_pow5}),
walther@60260
   742
	           (*"(a + b) \<up> 5 = ... "*)
walther@59871
   743
	         Rule.Thm ("real_plus_binom_pow3",ThmC.numerals_to_Free @{thm real_plus_binom_pow3}),
walther@60260
   744
	           (*"(a + b) \<up> 3 = a \<up> 3 + 3*a \<up> 2*b + 3*a*b \<up> 2 + b \<up> 3" *)
neuper@42407
   745
	         (*WN071229 changed/removed for Schaerding -----vvv*)
walther@59871
   746
	         (*Rule.Thm ("real_plus_binom_pow2",ThmC.numerals_to_Free @{thm real_plus_binom_pow2}),*)
walther@60260
   747
	           (*"(a + b) \<up> 2 = a \<up> 2 + 2*a*b + b \<up> 2"*)
walther@59871
   748
	         Rule.Thm ("real_plus_binom_pow2",ThmC.numerals_to_Free @{thm real_plus_binom_pow2}),
walther@60260
   749
	           (*"(a + b) \<up> 2 = (a + b) * (a + b)"*)
walther@59871
   750
	         (*Rule.Thm ("real_plus_minus_binom1_p_p", ThmC.numerals_to_Free @{thm real_plus_minus_binom1_p_p}),*)
walther@60260
   751
	           (*"(a + b)*(a + -1 * b) = a \<up> 2 + -1*b \<up> 2"*)
walther@59871
   752
	         (*Rule.Thm ("real_plus_minus_binom2_p_p", ThmC.numerals_to_Free @{thm real_plus_minus_binom2_p_p}),*)
walther@60260
   753
	           (*"(a + -1 * b)*(a + b) = a \<up> 2 + -1*b \<up> 2"*)
walther@60242
   754
	         (*WN071229 changed/removed for Schaerding -----\<up>*)
neuper@37950
   755
	      
walther@59871
   756
	         Rule.Thm ("distrib_right" ,ThmC.numerals_to_Free @{thm distrib_right}),
neuper@42407
   757
	           (*"(z1.0 + z2.0) * w = z1.0 * w + z2.0 * w"*)
walther@59871
   758
	         Rule.Thm ("distrib_left",ThmC.numerals_to_Free @{thm distrib_left}),
neuper@42407
   759
	           (*"w * (z1.0 + z2.0) = w * z1.0 + w * z2.0"*)
neuper@37950
   760
	       
walther@59871
   761
	         Rule.Thm ("realpow_multI", ThmC.numerals_to_Free @{thm realpow_multI}),
walther@60242
   762
	           (*"(r * s) \<up> n = r \<up> n * s \<up> n"*)
walther@60318
   763
	         Rule.Thm ("realpow_pow",ThmC.numerals_to_Free @{thm realpow_pow}),
walther@60242
   764
	           (*"(a \<up> b) \<up> c = a \<up> (b * c)"*)
walther@60318
   765
	         Rule.Thm ("realpow_minus_even",ThmC.numerals_to_Free @{thm realpow_minus_even}),
walther@60318
   766
	           (*"n is_even ==> (- r) \<up> n = r \<up> n"*)
walther@60318
   767
	         Rule.Thm ("realpow_minus_odd",ThmC.numerals_to_Free @{thm realpow_minus_odd})
walther@60318
   768
	           (*"Not (n is_even) ==> (- r) \<up> n = -1 * r \<up> n"*)
walther@59878
   769
	       ], scr = Rule.Empty_Prog};
neuper@37950
   770
neuper@37950
   771
val expand_poly_rat_ = 
walther@59851
   772
  Rule_Def.Repeat{id = "expand_poly_rat_", preconds = [], 
walther@59857
   773
      rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
walther@59852
   774
      erls =  Rule_Set.append_rules "Rule_Set.empty-is_polyexp" Rule_Set.empty
walther@60278
   775
	        [Rule.Eval ("Poly.is_polyexp", eval_is_polyexp "")
neuper@37950
   776
		 ],
walther@59851
   777
      srls = Rule_Set.Empty,
neuper@42451
   778
      calc = [], errpatts = [],
neuper@37950
   779
      rules = 
walther@59871
   780
        [Rule.Thm ("real_plus_binom_pow4_poly", ThmC.numerals_to_Free @{thm real_plus_binom_pow4_poly}),
walther@60260
   781
	     (*"[| a is_polyexp; b is_polyexp |] ==> (a + b) \<up> 4 = ... "*)
walther@59871
   782
	 Rule.Thm ("real_plus_binom_pow5_poly", ThmC.numerals_to_Free @{thm real_plus_binom_pow5_poly}),
walther@60260
   783
	     (*"[| a is_polyexp; b is_polyexp |] ==> (a + b) \<up> 5 = ... "*)
walther@59871
   784
	 Rule.Thm ("real_plus_binom_pow2_poly",ThmC.numerals_to_Free @{thm real_plus_binom_pow2_poly}),
neuper@37950
   785
	     (*"[| a is_polyexp; b is_polyexp |] ==>
walther@60260
   786
		            (a + b) \<up> 2 = a \<up> 2 + 2*a*b + b \<up> 2"*)
walther@59871
   787
	 Rule.Thm ("real_plus_binom_pow3_poly",ThmC.numerals_to_Free @{thm real_plus_binom_pow3_poly}),
neuper@37950
   788
	     (*"[| a is_polyexp; b is_polyexp |] ==> 
walther@60260
   789
			(a + b) \<up> 3 = a \<up> 3 + 3*a \<up> 2*b + 3*a*b \<up> 2 + b \<up> 3" *)
walther@59871
   790
	 Rule.Thm ("real_plus_minus_binom1_p_p",ThmC.numerals_to_Free @{thm real_plus_minus_binom1_p_p}),
walther@60260
   791
	     (*"(a + b)*(a + -1 * b) = a \<up> 2 + -1*b \<up> 2"*)
walther@59871
   792
	 Rule.Thm ("real_plus_minus_binom2_p_p",ThmC.numerals_to_Free @{thm real_plus_minus_binom2_p_p}),
walther@60260
   793
	     (*"(a + -1 * b)*(a + b) = a \<up> 2 + -1*b \<up> 2"*)
neuper@37950
   794
	      
wneuper@59416
   795
	 Rule.Thm ("real_add_mult_distrib_poly",
walther@59871
   796
               ThmC.numerals_to_Free @{thm real_add_mult_distrib_poly}),
neuper@37950
   797
	       (*"w is_polyexp ==> (z1.0 + z2.0) * w = z1.0 * w + z2.0 * w"*)
wneuper@59416
   798
	 Rule.Thm("real_add_mult_distrib2_poly",
walther@59871
   799
              ThmC.numerals_to_Free @{thm real_add_mult_distrib2_poly}),
neuper@37950
   800
	     (*"w is_polyexp ==> w * (z1.0 + z2.0) = w * z1.0 + w * z2.0"*)
neuper@37950
   801
	       
walther@59871
   802
	 Rule.Thm ("realpow_multI_poly", ThmC.numerals_to_Free @{thm realpow_multI_poly}),
neuper@37950
   803
	     (*"[| r is_polyexp; s is_polyexp |] ==> 
walther@60242
   804
		            (r * s) \<up> n = r \<up> n * s \<up> n"*)
walther@60318
   805
	 Rule.Thm ("realpow_pow",ThmC.numerals_to_Free @{thm realpow_pow}),
walther@60318
   806
	   (*"(a \<up> b) \<up> c = a \<up> (b * c)"*)
walther@60318
   807
	 Rule.Thm ("realpow_minus_even",ThmC.numerals_to_Free @{thm realpow_minus_even}),
walther@60318
   808
	   (*"n is_even ==> (- r) \<up> n = r \<up> n"*)
walther@60318
   809
	 Rule.Thm ("realpow_minus_odd",ThmC.numerals_to_Free @{thm realpow_minus_odd})
walther@60318
   810
	   (*"Not (n is_even) ==> (- r) \<up> n = -1 * r \<up> n"*)
walther@59878
   811
	 ], scr = Rule.Empty_Prog};
neuper@37950
   812
neuper@37950
   813
val simplify_power_ = 
walther@59851
   814
  Rule_Def.Repeat{id = "simplify_power_", preconds = [], 
walther@59857
   815
      rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
walther@59852
   816
      erls = Rule_Set.empty, srls = Rule_Set.Empty,
neuper@42451
   817
      calc = [], errpatts = [],
wneuper@59416
   818
      rules = [(*MG: Reihenfolge der folgenden 2 Rule.Thm muss so bleiben, wegen
walther@60260
   819
		a*(a*a) --> a*a \<up> 2 und nicht a*(a*a) --> a \<up> 2*a *)
wneuper@59416
   820
	       Rule.Thm ("sym_realpow_twoI",
walther@59871
   821
                     ThmC.numerals_to_Free (@{thm realpow_twoI} RS @{thm sym})),	
walther@60242
   822
	       (*"r * r = r \<up> 2"*)
walther@59871
   823
	       Rule.Thm ("realpow_twoI_assoc_l",ThmC.numerals_to_Free @{thm realpow_twoI_assoc_l}),
walther@60242
   824
	       (*"r * (r * s) = r \<up> 2 * s"*)
neuper@37950
   825
walther@59871
   826
	       Rule.Thm ("realpow_plus_1",ThmC.numerals_to_Free @{thm realpow_plus_1}),		
walther@60242
   827
	       (*"r * r \<up> n = r \<up> (n + 1)"*)
wneuper@59416
   828
	       Rule.Thm ("realpow_plus_1_assoc_l",
walther@59871
   829
                     ThmC.numerals_to_Free @{thm realpow_plus_1_assoc_l}),
walther@60242
   830
	       (*"r * (r \<up> m * s) = r \<up> (1 + m) * s"*)
walther@60260
   831
	       (*MG 9.7.03: neues Rule.Thm wegen a*(a*(a*b)) --> a \<up> 2*(a*b) *)
wneuper@59416
   832
	       Rule.Thm ("realpow_plus_1_assoc_l2",
walther@59871
   833
                     ThmC.numerals_to_Free @{thm realpow_plus_1_assoc_l2}),
walther@60242
   834
	       (*"r \<up> m * (r * s) = r \<up> (1 + m) * s"*)
neuper@37950
   835
wneuper@59416
   836
	       Rule.Thm ("sym_realpow_addI",
walther@59871
   837
               ThmC.numerals_to_Free (@{thm realpow_addI} RS @{thm sym})),
walther@60242
   838
	       (*"r \<up> n * r \<up> m = r \<up> (n + m)"*)
walther@59871
   839
	       Rule.Thm ("realpow_addI_assoc_l",ThmC.numerals_to_Free @{thm realpow_addI_assoc_l}),
walther@60242
   840
	       (*"r \<up> n * (r \<up> m * s) = r \<up> (n + m) * s"*)
neuper@37950
   841
	       
neuper@37950
   842
	       (* ist in expand_poly - wird hier aber auch gebraucht, wegen: 
walther@60260
   843
		  "r * r = r \<up> 2" wenn r=a \<up> b*)
walther@59871
   844
	       Rule.Thm ("realpow_pow",ThmC.numerals_to_Free @{thm realpow_pow})
walther@60242
   845
	       (*"(a \<up> b) \<up> c = a \<up> (b * c)"*)
walther@59878
   846
	       ], scr = Rule.Empty_Prog};
neuper@37950
   847
neuper@37950
   848
val calc_add_mult_pow_ = 
walther@59851
   849
  Rule_Def.Repeat{id = "calc_add_mult_pow_", preconds = [], 
walther@59857
   850
      rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
walther@59851
   851
      erls = Atools_erls(*erls3.4.03*),srls = Rule_Set.Empty,
walther@60278
   852
      calc = [("PLUS"  , ("Groups.plus_class.plus", eval_binop "#add_")), 
walther@60278
   853
	      ("TIMES" , ("Groups.times_class.times", eval_binop "#mult_")),
walther@60278
   854
	      ("POWER", ("Transcendental.powr", eval_binop "#power_"))
neuper@37950
   855
	      ],
neuper@42451
   856
      errpatts = [],
walther@60278
   857
      rules = [Rule.Eval ("Groups.plus_class.plus", eval_binop "#add_"),
walther@60278
   858
	       Rule.Eval ("Groups.times_class.times", eval_binop "#mult_"),
walther@60278
   859
	       Rule.Eval ("Transcendental.powr", eval_binop "#power_")
walther@59878
   860
	       ], scr = Rule.Empty_Prog};
neuper@37950
   861
neuper@37950
   862
val reduce_012_mult_ = 
walther@59851
   863
  Rule_Def.Repeat{id = "reduce_012_mult_", preconds = [], 
walther@59857
   864
      rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
walther@59852
   865
      erls = Rule_Set.empty,srls = Rule_Set.Empty,
neuper@42451
   866
      calc = [], errpatts = [],
wneuper@59416
   867
      rules = [(* MG: folgende Rule.Thm müssen hier stehen bleiben: *)
walther@59871
   868
               Rule.Thm ("mult_1_right",ThmC.numerals_to_Free @{thm mult_1_right}),
walther@60260
   869
	       (*"z * 1 = z"*) (*wegen "a * b * b \<up> (-1) + a"*) 
walther@59871
   870
	       Rule.Thm ("realpow_zeroI",ThmC.numerals_to_Free @{thm realpow_zeroI}),
walther@60260
   871
	       (*"r \<up> 0 = 1"*) (*wegen "a*a \<up> (-1)*c + b + c"*)
walther@59871
   872
	       Rule.Thm ("realpow_oneI",ThmC.numerals_to_Free @{thm realpow_oneI}),
walther@60242
   873
	       (*"r \<up> 1 = r"*)
walther@59871
   874
	       Rule.Thm ("realpow_eq_oneI",ThmC.numerals_to_Free @{thm realpow_eq_oneI})
walther@60242
   875
	       (*"1 \<up> n = 1"*)
walther@59878
   876
	       ], scr = Rule.Empty_Prog};
neuper@37950
   877
neuper@37950
   878
val collect_numerals_ = 
walther@59851
   879
  Rule_Def.Repeat{id = "collect_numerals_", preconds = [], 
walther@59857
   880
      rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
walther@59851
   881
      erls = Atools_erls, srls = Rule_Set.Empty,
walther@60278
   882
      calc = [("PLUS"  , ("Groups.plus_class.plus", eval_binop "#add_"))
neuper@42451
   883
	      ], errpatts = [],
neuper@37950
   884
      rules = 
walther@59871
   885
        [Rule.Thm ("real_num_collect",ThmC.numerals_to_Free @{thm real_num_collect}), 
neuper@37950
   886
	     (*"[| l is_const; m is_const |]==>l * n + m * n = (l + m) * n"*)
walther@59871
   887
	 Rule.Thm ("real_num_collect_assoc_r",ThmC.numerals_to_Free @{thm real_num_collect_assoc_r}),
neuper@37950
   888
	     (*"[| l is_const; m is_const |] ==>  \
neuper@37950
   889
					\(k + m * n) + l * n = k + (l + m)*n"*)
walther@59871
   890
	 Rule.Thm ("real_one_collect",ThmC.numerals_to_Free @{thm real_one_collect}),	
neuper@37950
   891
	     (*"m is_const ==> n + m * n = (1 + m) * n"*)
walther@59871
   892
	 Rule.Thm ("real_one_collect_assoc_r",ThmC.numerals_to_Free @{thm real_one_collect_assoc_r}), 
neuper@37950
   893
	     (*"m is_const ==> (k + n) + m * n = k + (m + 1) * n"*)
neuper@37950
   894
walther@60278
   895
         Rule.Eval ("Groups.plus_class.plus", eval_binop "#add_"),
neuper@37950
   896
wneuper@59416
   897
	 (*MG: Reihenfolge der folgenden 2 Rule.Thm muss so bleiben, wegen
neuper@37950
   898
		     (a+a)+a --> a + 2*a --> 3*a and not (a+a)+a --> 2*a + a *)
walther@59871
   899
         Rule.Thm ("real_mult_2_assoc_r",ThmC.numerals_to_Free @{thm real_mult_2_assoc_r}),
neuper@37950
   900
	     (*"(k + z1) + z1 = k + 2 * z1"*)
walther@59871
   901
	 Rule.Thm ("sym_real_mult_2",ThmC.numerals_to_Free (@{thm real_mult_2} RS @{thm sym}))
neuper@37950
   902
	     (*"z1 + z1 = 2 * z1"*)
walther@59878
   903
	], scr = Rule.Empty_Prog};
neuper@37950
   904
neuper@37950
   905
val reduce_012_ = 
walther@59851
   906
  Rule_Def.Repeat{id = "reduce_012_", preconds = [], 
walther@59857
   907
      rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
walther@59852
   908
      erls = Rule_Set.empty,srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@59871
   909
      rules = [Rule.Thm ("mult_1_left",ThmC.numerals_to_Free @{thm mult_1_left}),                 
neuper@37950
   910
	       (*"1 * z = z"*)
walther@59871
   911
	       Rule.Thm ("mult_zero_left",ThmC.numerals_to_Free @{thm mult_zero_left}),        
neuper@37950
   912
	       (*"0 * z = 0"*)
walther@59871
   913
	       Rule.Thm ("mult_zero_right",ThmC.numerals_to_Free @{thm mult_zero_right}),
neuper@37950
   914
	       (*"z * 0 = 0"*)
walther@59871
   915
	       Rule.Thm ("add_0_left",ThmC.numerals_to_Free @{thm add_0_left}),
neuper@37950
   916
	       (*"0 + z = z"*)
walther@59871
   917
	       Rule.Thm ("add_0_right",ThmC.numerals_to_Free @{thm add_0_right}),
neuper@37950
   918
	       (*"z + 0 = z"*) (*wegen a+b-b --> a+(1-1)*b --> a+0 --> a*)
neuper@37950
   919
walther@59871
   920
	       (*Rule.Thm ("realpow_oneI",ThmC.numerals_to_Free @{thm realpow_oneI})*)
walther@60242
   921
	       (*"?r \<up> 1 = ?r"*)
walther@59871
   922
	       Rule.Thm ("division_ring_divide_zero",ThmC.numerals_to_Free @{thm division_ring_divide_zero})
neuper@37950
   923
	       (*"0 / ?x = 0"*)
walther@59878
   924
	       ], scr = Rule.Empty_Prog};
neuper@37950
   925
neuper@37979
   926
val discard_parentheses1 = 
walther@59852
   927
    Rule_Set.append_rules "discard_parentheses1" Rule_Set.empty 
walther@59877
   928
	       [Rule.Thm ("sym_mult.assoc",
walther@59871
   929
                      ThmC.numerals_to_Free (@{thm mult.assoc} RS @{thm sym}))
neuper@37950
   930
		(*"?z1.1 * (?z2.1 * ?z3.1) = ?z1.1 * ?z2.1 * ?z3.1"*)
walther@59877
   931
		(*Rule.Thm ("sym_add.assoc",
walther@59877
   932
                        ThmC.numerals_to_Free (@{thm add.assoc} RS @{thm sym}))*)
neuper@37950
   933
		(*"?z1.1 + (?z2.1 + ?z3.1) = ?z1.1 + ?z2.1 + ?z3.1"*)
neuper@37950
   934
		 ];
neuper@37950
   935
wneuper@59523
   936
val expand_poly =
walther@59851
   937
  Rule_Def.Repeat{id = "expand_poly", preconds = [], 
walther@59857
   938
      rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
walther@60320
   939
      erls = powers_erls, srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@60320
   940
      rules = 
walther@60320
   941
        [Rule.Thm ("distrib_right" ,ThmC.numerals_to_Free @{thm distrib_right}),
neuper@37950
   942
	       (*"(z1.0 + z2.0) * w = z1.0 * w + z2.0 * w"*)
walther@59871
   943
	       Rule.Thm ("distrib_left",ThmC.numerals_to_Free @{thm distrib_left}),
neuper@37950
   944
	       (*"w * (z1.0 + z2.0) = w * z1.0 + w * z2.0"*)
walther@59871
   945
	       (*Rule.Thm ("distrib_right1",ThmC.numerals_to_Free @{thm distrib_right}1),
neuper@37950
   946
		....... 18.3.03 undefined???*)
neuper@37950
   947
walther@59871
   948
	       Rule.Thm ("real_plus_binom_pow2",ThmC.numerals_to_Free @{thm real_plus_binom_pow2}),
walther@60260
   949
	       (*"(a + b) \<up> 2 = a \<up> 2 + 2*a*b + b \<up> 2"*)
walther@59871
   950
	       Rule.Thm ("real_minus_binom_pow2_p",ThmC.numerals_to_Free @{thm real_minus_binom_pow2_p}),
walther@60260
   951
	       (*"(a - b) \<up> 2 = a \<up> 2 + -2*a*b + b \<up> 2"*)
wneuper@59416
   952
	       Rule.Thm ("real_plus_minus_binom1_p",
walther@59871
   953
		    ThmC.numerals_to_Free @{thm real_plus_minus_binom1_p}),
walther@60260
   954
	       (*"(a + b)*(a - b) = a \<up> 2 + -1*b \<up> 2"*)
wneuper@59416
   955
	       Rule.Thm ("real_plus_minus_binom2_p",
walther@59871
   956
		    ThmC.numerals_to_Free @{thm real_plus_minus_binom2_p}),
walther@60260
   957
	       (*"(a - b)*(a + b) = a \<up> 2 + -1*b \<up> 2"*)
neuper@37950
   958
walther@59871
   959
	       Rule.Thm ("minus_minus",ThmC.numerals_to_Free @{thm minus_minus}),
neuper@37950
   960
	       (*"- (- ?z) = ?z"*)
walther@59871
   961
	       Rule.Thm ("real_diff_minus",ThmC.numerals_to_Free @{thm real_diff_minus}),
neuper@37950
   962
	       (*"a - b = a + -1 * b"*)
walther@60320
   963
	       Rule.Thm ("real_mult_minus1_sym", ThmC.numerals_to_Free (@{thm real_mult_minus1_sym}))
walther@60320
   964
	       (*"\<not>(z is_const) ==> - (z::real) = -1 * z"*)
neuper@37950
   965
wneuper@59416
   966
	       (*Rule.Thm ("real_minus_add_distrib",
walther@59871
   967
		      ThmC.numerals_to_Free @{thm real_minus_add_distrib}),*)
neuper@37950
   968
	       (*"- (?x + ?y) = - ?x + - ?y"*)
walther@59871
   969
	       (*Rule.Thm ("real_diff_plus",ThmC.numerals_to_Free @{thm real_diff_plus})*)
neuper@37950
   970
	       (*"a - b = a + -b"*)
walther@59878
   971
	       ], scr = Rule.Empty_Prog};
neuper@37950
   972
neuper@37950
   973
val simplify_power = 
walther@59851
   974
  Rule_Def.Repeat{id = "simplify_power", preconds = [], 
walther@59857
   975
      rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
walther@59852
   976
      erls = Rule_Set.empty, srls = Rule_Set.Empty,
neuper@42451
   977
      calc = [], errpatts = [],
walther@59871
   978
      rules = [Rule.Thm ("realpow_multI", ThmC.numerals_to_Free @{thm realpow_multI}),
walther@60242
   979
	       (*"(r * s) \<up> n = r \<up> n * s \<up> n"*)
neuper@37950
   980
	       
wneuper@59416
   981
	       Rule.Thm ("sym_realpow_twoI",
walther@59871
   982
                     ThmC.numerals_to_Free( @{thm realpow_twoI} RS @{thm sym})),	
walther@60242
   983
	       (*"r1 * r1 = r1 \<up> 2"*)
walther@59871
   984
	       Rule.Thm ("realpow_plus_1",ThmC.numerals_to_Free @{thm realpow_plus_1}),		
walther@60242
   985
	       (*"r * r \<up> n = r \<up> (n + 1)"*)
walther@59871
   986
	       Rule.Thm ("realpow_pow",ThmC.numerals_to_Free @{thm realpow_pow}),
walther@60242
   987
	       (*"(a \<up> b) \<up> c = a \<up> (b * c)"*)
wneuper@59416
   988
	       Rule.Thm ("sym_realpow_addI",
walther@59871
   989
                     ThmC.numerals_to_Free (@{thm realpow_addI} RS @{thm sym})),
walther@60242
   990
	       (*"r \<up> n * r \<up> m = r \<up> (n + m)"*)
walther@59871
   991
	       Rule.Thm ("realpow_oneI",ThmC.numerals_to_Free @{thm realpow_oneI}),
walther@60242
   992
	       (*"r \<up> 1 = r"*)
walther@59871
   993
	       Rule.Thm ("realpow_eq_oneI",ThmC.numerals_to_Free @{thm realpow_eq_oneI})
walther@60242
   994
	       (*"1 \<up> n = 1"*)
walther@59878
   995
	       ], scr = Rule.Empty_Prog};
neuper@42451
   996
neuper@37950
   997
val collect_numerals = 
walther@59851
   998
  Rule_Def.Repeat{id = "collect_numerals", preconds = [], 
walther@59857
   999
      rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
walther@59851
  1000
      erls = Atools_erls(*erls3.4.03*),srls = Rule_Set.Empty,
walther@60278
  1001
      calc = [("PLUS"  , ("Groups.plus_class.plus", eval_binop "#add_")), 
walther@60278
  1002
	      ("TIMES" , ("Groups.times_class.times", eval_binop "#mult_")),
walther@60278
  1003
	      ("POWER", ("Transcendental.powr", eval_binop "#power_"))
neuper@42451
  1004
	      ], errpatts = [],
walther@59871
  1005
      rules = [Rule.Thm ("real_num_collect",ThmC.numerals_to_Free @{thm real_num_collect}), 
neuper@37950
  1006
	       (*"[| l is_const; m is_const |]==>l * n + m * n = (l + m) * n"*)
walther@59871
  1007
	       Rule.Thm ("real_num_collect_assoc",ThmC.numerals_to_Free @{thm real_num_collect_assoc}),
neuper@37950
  1008
	       (*"[| l is_const; m is_const |] ==>  
neuper@37950
  1009
				l * n + (m * n + k) =  (l + m) * n + k"*)
walther@59871
  1010
	       Rule.Thm ("real_one_collect",ThmC.numerals_to_Free @{thm real_one_collect}),	
neuper@37950
  1011
	       (*"m is_const ==> n + m * n = (1 + m) * n"*)
walther@59871
  1012
	       Rule.Thm ("real_one_collect_assoc",ThmC.numerals_to_Free @{thm real_one_collect_assoc}), 
neuper@37950
  1013
	       (*"m is_const ==> k + (n + m * n) = k + (1 + m) * n"*)
walther@60278
  1014
	       Rule.Eval ("Groups.plus_class.plus", eval_binop "#add_"), 
walther@60278
  1015
	       Rule.Eval ("Groups.times_class.times", eval_binop "#mult_"),
walther@60278
  1016
	       Rule.Eval ("Transcendental.powr", eval_binop "#power_")
walther@59878
  1017
	       ], scr = Rule.Empty_Prog};
neuper@37950
  1018
val reduce_012 = 
walther@59851
  1019
  Rule_Def.Repeat{id = "reduce_012", preconds = [], 
walther@59857
  1020
      rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
walther@59852
  1021
      erls = Rule_Set.empty,srls = Rule_Set.Empty,
neuper@42451
  1022
      calc = [], errpatts = [],
walther@59871
  1023
      rules = [Rule.Thm ("mult_1_left",ThmC.numerals_to_Free @{thm mult_1_left}),                 
neuper@37950
  1024
	       (*"1 * z = z"*)
walther@59871
  1025
	       (*Rule.Thm ("real_mult_minus1",ThmC.numerals_to_Free @{thm real_mult_minus1}),14.3.03*)
neuper@37950
  1026
	       (*"-1 * z = - z"*)
wneuper@59416
  1027
	       Rule.Thm ("minus_mult_left", 
walther@59871
  1028
		    ThmC.numerals_to_Free (@{thm minus_mult_left} RS @{thm sym})),
neuper@37950
  1029
	       (*- (?x * ?y) = "- ?x * ?y"*)
wneuper@59416
  1030
	       (*Rule.Thm ("real_minus_mult_cancel",
walther@59871
  1031
                       ThmC.numerals_to_Free @{thm real_minus_mult_cancel}),
neuper@37950
  1032
	       (*"- ?x * - ?y = ?x * ?y"*)---*)
walther@59871
  1033
	       Rule.Thm ("mult_zero_left",ThmC.numerals_to_Free @{thm mult_zero_left}),        
neuper@37950
  1034
	       (*"0 * z = 0"*)
walther@59871
  1035
	       Rule.Thm ("add_0_left",ThmC.numerals_to_Free @{thm add_0_left}),
neuper@37950
  1036
	       (*"0 + z = z"*)
walther@59871
  1037
	       Rule.Thm ("right_minus",ThmC.numerals_to_Free @{thm right_minus}),
neuper@37950
  1038
	       (*"?z + - ?z = 0"*)
wneuper@59416
  1039
	       Rule.Thm ("sym_real_mult_2",
walther@59871
  1040
                     ThmC.numerals_to_Free (@{thm real_mult_2} RS @{thm sym})),	
neuper@37950
  1041
	       (*"z1 + z1 = 2 * z1"*)
walther@59871
  1042
	       Rule.Thm ("real_mult_2_assoc",ThmC.numerals_to_Free @{thm real_mult_2_assoc})
neuper@37950
  1043
	       (*"z1 + (z1 + k) = 2 * z1 + k"*)
walther@59878
  1044
	       ], scr = Rule.Empty_Prog};
neuper@52139
  1045
neuper@37950
  1046
val discard_parentheses = 
walther@59852
  1047
    Rule_Set.append_rules "discard_parentheses" Rule_Set.empty 
walther@59877
  1048
	       [Rule.Thm ("sym_mult.assoc",
walther@59871
  1049
                      ThmC.numerals_to_Free (@{thm mult.assoc} RS @{thm sym})),
walther@59877
  1050
		Rule.Thm ("sym_add.assoc",
walther@59871
  1051
                      ThmC.numerals_to_Free (@{thm add.assoc} RS @{thm sym}))];
wneuper@59523
  1052
\<close>
neuper@37950
  1053
wneuper@59523
  1054
subsubsection \<open>hard-coded AC rewriting\<close>
wneuper@59523
  1055
ML \<open>
wneuper@59523
  1056
(*MG.0401: termorders for multivariate polys dropped due to principal problems:
wneuper@59523
  1057
  (total-degree-)ordering of monoms NOT possible with size_of_term GIVEN*)
wneuper@59523
  1058
val order_add_mult = 
walther@59851
  1059
  Rule_Def.Repeat{id = "order_add_mult", preconds = [], 
wneuper@59523
  1060
      rew_ord = ("ord_make_polynomial",ord_make_polynomial false thy),
walther@59852
  1061
      erls = Rule_Set.empty,srls = Rule_Set.Empty,
neuper@42451
  1062
      calc = [], errpatts = [],
walther@59877
  1063
      rules = [Rule.Thm ("mult.commute",ThmC.numerals_to_Free @{thm mult.commute}),
wneuper@59523
  1064
	       (* z * w = w * z *)
walther@59871
  1065
	       Rule.Thm ("real_mult_left_commute",ThmC.numerals_to_Free @{thm real_mult_left_commute}),
wneuper@59523
  1066
	       (*z1.0 * (z2.0 * z3.0) = z2.0 * (z1.0 * z3.0)*)
walther@59877
  1067
	       Rule.Thm ("mult.assoc",ThmC.numerals_to_Free @{thm mult.assoc}),		
wneuper@59523
  1068
	       (*z1.0 * z2.0 * z3.0 = z1.0 * (z2.0 * z3.0)*)
walther@59877
  1069
	       Rule.Thm ("add.commute",ThmC.numerals_to_Free @{thm add.commute}),	
wneuper@59523
  1070
	       (*z + w = w + z*)
walther@59877
  1071
	       Rule.Thm ("add.left_commute",ThmC.numerals_to_Free @{thm add.left_commute}),
wneuper@59523
  1072
	       (*x + (y + z) = y + (x + z)*)
walther@59877
  1073
	       Rule.Thm ("add.assoc",ThmC.numerals_to_Free @{thm add.assoc})	               
wneuper@59523
  1074
	       (*z1.0 + z2.0 + z3.0 = z1.0 + (z2.0 + z3.0)*)
walther@59878
  1075
	       ], scr = Rule.Empty_Prog};
wneuper@59523
  1076
(*MG.0401: termorders for multivariate polys dropped due to principal problems:
wneuper@59523
  1077
  (total-degree-)ordering of monoms NOT possible with size_of_term GIVEN*)
wneuper@59523
  1078
val order_mult = 
walther@59851
  1079
  Rule_Def.Repeat{id = "order_mult", preconds = [], 
wneuper@59523
  1080
      rew_ord = ("ord_make_polynomial",ord_make_polynomial false thy),
walther@59852
  1081
      erls = Rule_Set.empty,srls = Rule_Set.Empty,
wneuper@59523
  1082
      calc = [], errpatts = [],
walther@59877
  1083
      rules = [Rule.Thm ("mult.commute",ThmC.numerals_to_Free @{thm mult.commute}),
wneuper@59523
  1084
	       (* z * w = w * z *)
walther@59871
  1085
	       Rule.Thm ("real_mult_left_commute",ThmC.numerals_to_Free @{thm real_mult_left_commute}),
wneuper@59523
  1086
	       (*z1.0 * (z2.0 * z3.0) = z2.0 * (z1.0 * z3.0)*)
walther@59877
  1087
	       Rule.Thm ("mult.assoc",ThmC.numerals_to_Free @{thm mult.assoc})	
wneuper@59523
  1088
	       (*z1.0 * z2.0 * z3.0 = z1.0 * (z2.0 * z3.0)*)
walther@59878
  1089
	       ], scr = Rule.Empty_Prog};
wneuper@59472
  1090
\<close>
wneuper@59472
  1091
ML \<open>
wneuper@59416
  1092
fun attach_form (_: Rule.rule list list) (_: term) (_: term) = (*still missing*)
wneuper@59416
  1093
    []:(Rule.rule * (term * term list)) list;
walther@59850
  1094
fun init_state (_: term) = Rule_Set.e_rrlsstate;
wneuper@59416
  1095
fun locate_rule (_: Rule.rule list list) (_: term) (_: Rule.rule) =
wneuper@59416
  1096
    ([]:(Rule.rule * (term * term list)) list);
wneuper@59416
  1097
fun next_rule (_: Rule.rule list list) (_: term) = (NONE: Rule.rule option);
wneuper@59406
  1098
fun normal_form t = SOME (sort_variables t, []: term list);
neuper@37950
  1099
neuper@37950
  1100
val order_mult_ =
walther@59850
  1101
    Rule_Set.Rrls {id = "order_mult_", 
neuper@37950
  1102
	  prepat = 
neuper@38036
  1103
          (* ?p matched with the current term gives an environment,
neuper@38037
  1104
             which evaluates (the instantiated) "?p is_multUnordered" to true *)
wneuper@59389
  1105
	  [([TermC.parse_patt thy "?p is_multUnordered"], 
wneuper@59389
  1106
             TermC.parse_patt thy "?p :: real")],
walther@59857
  1107
	  rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
walther@59852
  1108
	  erls = Rule_Set.append_rules "Rule_Set.empty-is_multUnordered" Rule_Set.empty
walther@60278
  1109
			    [Rule.Eval ("Poly.is_multUnordered", 
neuper@37976
  1110
                                    eval_is_multUnordered "")],
walther@60278
  1111
	  calc = [("PLUS"  , ("Groups.plus_class.plus", eval_binop "#add_")),
walther@60278
  1112
		  ("TIMES" , ("Groups.times_class.times", eval_binop "#mult_")),
walther@59603
  1113
		  ("DIVIDE", ("Rings.divide_class.divide", Prog_Expr.eval_cancel "#divide_e")),
walther@60278
  1114
		  ("POWER" , ("Transcendental.powr", eval_binop "#power_"))],
wneuper@59406
  1115
    errpatts = [],
wneuper@59416
  1116
	  scr = Rule.Rfuns {init_state  = init_state,
neuper@37950
  1117
		     normal_form = normal_form,
neuper@37950
  1118
		     locate_rule = locate_rule,
neuper@37950
  1119
		     next_rule   = next_rule,
neuper@37950
  1120
		     attach_form = attach_form}};
neuper@37950
  1121
val order_mult_rls_ = 
walther@59851
  1122
  Rule_Def.Repeat {id = "order_mult_rls_", preconds = [], 
walther@59857
  1123
      rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
walther@59852
  1124
      erls = Rule_Set.empty,srls = Rule_Set.Empty,
neuper@42451
  1125
      calc = [], errpatts = [],
wneuper@59416
  1126
      rules = [Rule.Rls_ order_mult_
walther@59878
  1127
	       ], scr = Rule.Empty_Prog};
neuper@37950
  1128
wneuper@59523
  1129
\<close> ML \<open>
neuper@37950
  1130
wneuper@59416
  1131
fun attach_form (_: Rule.rule list list) (_: term) (_: term) = (*still missing*)
wneuper@59416
  1132
    []: (Rule.rule * (term * term list)) list;
walther@59850
  1133
fun init_state (_: term) = Rule_Set.e_rrlsstate;
wneuper@59416
  1134
fun locate_rule (_: Rule.rule list list) (_: term) (_: Rule.rule) =
wneuper@59416
  1135
    ([]: (Rule.rule * (term * term list)) list);
wneuper@59416
  1136
fun next_rule (_: Rule.rule list list) (_: term) = (NONE: Rule.rule option);
wneuper@59406
  1137
fun normal_form t = SOME (sort_monoms t,[]: term list);
wneuper@59472
  1138
\<close> ML \<open>
neuper@37950
  1139
val order_add_ =
walther@59850
  1140
    Rule_Set.Rrls {id = "order_add_", 
neuper@37950
  1141
	  prepat = (*WN.18.6.03 Preconditions und Pattern,
walther@59850
  1142
		    die beide passen muessen, damit das Rule_Set.Rrls angewandt wird*)
wneuper@59389
  1143
	  [([TermC.parse_patt @{theory} "?p is_addUnordered"], 
wneuper@59389
  1144
	     TermC.parse_patt @{theory} "?p :: real" 
neuper@37950
  1145
	    (*WN.18.6.03 also KEIN pattern, dieses erzeugt nur das Environment 
neuper@37950
  1146
	      fuer die Evaluation der Precondition "p is_addUnordered"*))],
walther@59857
  1147
	  rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
walther@59852
  1148
	  erls = Rule_Set.append_rules "Rule_Set.empty-is_addUnordered" Rule_Set.empty(*MG: poly_erls*)
walther@60278
  1149
			    [Rule.Eval ("Poly.is_addUnordered", eval_is_addUnordered "")],
walther@60278
  1150
	  calc = [("PLUS"  ,("Groups.plus_class.plus", eval_binop "#add_")),
walther@60278
  1151
		  ("TIMES" ,("Groups.times_class.times", eval_binop "#mult_")),
walther@59603
  1152
		  ("DIVIDE",("Rings.divide_class.divide", Prog_Expr.eval_cancel "#divide_e")),
walther@60278
  1153
		  ("POWER" ,("Transcendental.powr"  , eval_binop "#power_"))],
neuper@42451
  1154
	  errpatts = [],
wneuper@59416
  1155
	  scr = Rule.Rfuns {init_state  = init_state,
neuper@37950
  1156
		     normal_form = normal_form,
neuper@37950
  1157
		     locate_rule = locate_rule,
neuper@37950
  1158
		     next_rule   = next_rule,
neuper@37950
  1159
		     attach_form = attach_form}};
neuper@37950
  1160
wneuper@59406
  1161
val order_add_rls_ =
walther@59851
  1162
  Rule_Def.Repeat {id = "order_add_rls_", preconds = [], 
walther@59857
  1163
      rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
walther@59852
  1164
      erls = Rule_Set.empty,srls = Rule_Set.Empty,
neuper@42451
  1165
      calc = [], errpatts = [],
wneuper@59416
  1166
      rules = [Rule.Rls_ order_add_
walther@59878
  1167
	       ], scr = Rule.Empty_Prog};
wneuper@59472
  1168
\<close>
neuper@37950
  1169
wneuper@59472
  1170
text \<open>rule-set make_polynomial also named norm_Poly:
neuper@42398
  1171
  Rewrite order has not been implemented properly; the order is better in 
neuper@42398
  1172
  make_polynomial_in (coded in SML).
neuper@42398
  1173
  Notes on state of development:
neuper@42398
  1174
  \# surprise 2006: test --- norm_Poly NOT COMPLETE ---
neuper@42398
  1175
  \# migration Isabelle2002 --> 2011 weakened the rule set, see test
walther@59962
  1176
  --- Matthias Goldgruber 2003 rewrite orders ---, raise ERROR "ord_make_polynomial_in #16b"
wneuper@59472
  1177
\<close>
wneuper@59472
  1178
ML \<open>
neuper@37950
  1179
(*. see MG-DA.p.52ff .*)
neuper@37950
  1180
val make_polynomial(*MG.03, overwrites version from above, 
neuper@37950
  1181
    previously 'make_polynomial_'*) =
walther@59878
  1182
  Rule_Set.Sequence {id = "make_polynomial", preconds = []:term list, 
walther@59857
  1183
      rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
walther@59851
  1184
      erls = Atools_erls, srls = Rule_Set.Empty,calc = [], errpatts = [],
wneuper@59416
  1185
      rules = [Rule.Rls_ discard_minus,
wneuper@59416
  1186
	       Rule.Rls_ expand_poly_,
walther@60278
  1187
	       Rule.Eval ("Groups.times_class.times", eval_binop "#mult_"),
wneuper@59416
  1188
	       Rule.Rls_ order_mult_rls_,
wneuper@59416
  1189
	       Rule.Rls_ simplify_power_, 
wneuper@59416
  1190
	       Rule.Rls_ calc_add_mult_pow_, 
wneuper@59416
  1191
	       Rule.Rls_ reduce_012_mult_,
wneuper@59416
  1192
	       Rule.Rls_ order_add_rls_,
wneuper@59416
  1193
	       Rule.Rls_ collect_numerals_, 
wneuper@59416
  1194
	       Rule.Rls_ reduce_012_,
wneuper@59416
  1195
	       Rule.Rls_ discard_parentheses1
neuper@37950
  1196
	       ],
walther@59878
  1197
      scr = Rule.Empty_Prog
wneuper@59406
  1198
      };
wneuper@59472
  1199
\<close>
wneuper@59472
  1200
ML \<open>
neuper@37950
  1201
val norm_Poly(*=make_polynomial*) = 
walther@59878
  1202
  Rule_Set.Sequence {id = "norm_Poly", preconds = []:term list, 
walther@59857
  1203
      rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
walther@59851
  1204
      erls = Atools_erls, srls = Rule_Set.Empty, calc = [], errpatts = [],
wneuper@59416
  1205
      rules = [Rule.Rls_ discard_minus,
wneuper@59416
  1206
	       Rule.Rls_ expand_poly_,
walther@60278
  1207
	       Rule.Eval ("Groups.times_class.times", eval_binop "#mult_"),
wneuper@59416
  1208
	       Rule.Rls_ order_mult_rls_,
wneuper@59416
  1209
	       Rule.Rls_ simplify_power_, 
wneuper@59416
  1210
	       Rule.Rls_ calc_add_mult_pow_, 
wneuper@59416
  1211
	       Rule.Rls_ reduce_012_mult_,
wneuper@59416
  1212
	       Rule.Rls_ order_add_rls_,
wneuper@59416
  1213
	       Rule.Rls_ collect_numerals_, 
wneuper@59416
  1214
	       Rule.Rls_ reduce_012_,
wneuper@59416
  1215
	       Rule.Rls_ discard_parentheses1
neuper@37950
  1216
	       ],
walther@59878
  1217
      scr = Rule.Empty_Prog
wneuper@59406
  1218
      };
wneuper@59472
  1219
\<close>
wneuper@59472
  1220
ML \<open>
wneuper@59416
  1221
(* MG:03 Like make_polynomial_ but without Rule.Rls_ discard_parentheses1 
neuper@37950
  1222
   and expand_poly_rat_ instead of expand_poly_, see MG-DA.p.56ff*)
neuper@37950
  1223
(* MG necessary  for termination of norm_Rational(*_mg*) in Rational.ML*)
neuper@37950
  1224
val make_rat_poly_with_parentheses =
walther@59878
  1225
  Rule_Set.Sequence{id = "make_rat_poly_with_parentheses", preconds = []:term list, 
walther@59857
  1226
      rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
walther@59851
  1227
      erls = Atools_erls, srls = Rule_Set.Empty, calc = [], errpatts = [],
wneuper@59416
  1228
      rules = [Rule.Rls_ discard_minus,
wneuper@59416
  1229
	       Rule.Rls_ expand_poly_rat_,(*ignors rationals*)
walther@60278
  1230
	       Rule.Eval ("Groups.times_class.times", eval_binop "#mult_"),
wneuper@59416
  1231
	       Rule.Rls_ order_mult_rls_,
wneuper@59416
  1232
	       Rule.Rls_ simplify_power_, 
wneuper@59416
  1233
	       Rule.Rls_ calc_add_mult_pow_, 
wneuper@59416
  1234
	       Rule.Rls_ reduce_012_mult_,
wneuper@59416
  1235
	       Rule.Rls_ order_add_rls_,
wneuper@59416
  1236
	       Rule.Rls_ collect_numerals_, 
wneuper@59416
  1237
	       Rule.Rls_ reduce_012_
wneuper@59416
  1238
	       (*Rule.Rls_ discard_parentheses1 *)
neuper@37950
  1239
	       ],
walther@59878
  1240
      scr = Rule.Empty_Prog
wneuper@59406
  1241
      };
wneuper@59472
  1242
\<close>
wneuper@59472
  1243
ML \<open>
neuper@37950
  1244
(*.a minimal ruleset for reverse rewriting of factions [2];
neuper@37950
  1245
   compare expand_binoms.*)
neuper@37950
  1246
val rev_rew_p = 
walther@59878
  1247
Rule_Set.Sequence{id = "rev_rew_p", preconds = [], rew_ord = ("termlessI",termlessI),
walther@59851
  1248
    erls = Atools_erls, srls = Rule_Set.Empty,
walther@60278
  1249
    calc = [(*("PLUS"  , ("Groups.plus_class.plus", eval_binop "#add_")), 
walther@60278
  1250
	    ("TIMES" , ("Groups.times_class.times", eval_binop "#mult_")),
walther@60278
  1251
	    ("POWER", ("Transcendental.powr", eval_binop "#power_"))*)
neuper@42451
  1252
	    ], errpatts = [],
walther@59871
  1253
    rules = [Rule.Thm ("real_plus_binom_times" ,ThmC.numerals_to_Free @{thm real_plus_binom_times}),
neuper@37950
  1254
	     (*"(a + b)*(a + b) = a ^ 2 + 2 * a * b + b ^ 2*)
walther@59871
  1255
	     Rule.Thm ("real_plus_binom_times1" ,ThmC.numerals_to_Free @{thm real_plus_binom_times1}),
walther@60260
  1256
	     (*"(a +  1*b)*(a + -1*b) = a \<up> 2 + -1*b \<up> 2"*)
walther@59871
  1257
	     Rule.Thm ("real_plus_binom_times2" ,ThmC.numerals_to_Free @{thm real_plus_binom_times2}),
walther@60260
  1258
	     (*"(a + -1*b)*(a +  1*b) = a \<up> 2 + -1*b \<up> 2"*)
neuper@37950
  1259
walther@59871
  1260
	     Rule.Thm ("mult_1_left",ThmC.numerals_to_Free @{thm mult_1_left}),(*"1 * z = z"*)
neuper@37950
  1261
walther@59871
  1262
             Rule.Thm ("distrib_right" ,ThmC.numerals_to_Free @{thm distrib_right}),
neuper@37950
  1263
	     (*"(z1.0 + z2.0) * w = z1.0 * w + z2.0 * w"*)
walther@59871
  1264
	     Rule.Thm ("distrib_left",ThmC.numerals_to_Free @{thm distrib_left}),
neuper@37950
  1265
	     (*"w * (z1.0 + z2.0) = w * z1.0 + w * z2.0"*)
neuper@37950
  1266
	       
walther@59877
  1267
	     Rule.Thm ("mult.assoc", ThmC.numerals_to_Free @{thm mult.assoc}),
neuper@37950
  1268
	     (*"?z1.1 * ?z2.1 * ?z3. =1 ?z1.1 * (?z2.1 * ?z3.1)"*)
wneuper@59416
  1269
	     Rule.Rls_ order_mult_rls_,
wneuper@59416
  1270
	     (*Rule.Rls_ order_add_rls_,*)
neuper@37950
  1271
walther@60278
  1272
	     Rule.Eval ("Groups.plus_class.plus", eval_binop "#add_"), 
walther@60278
  1273
	     Rule.Eval ("Groups.times_class.times", eval_binop "#mult_"),
walther@60278
  1274
	     Rule.Eval ("Transcendental.powr", eval_binop "#power_"),
neuper@37950
  1275
	     
wneuper@59416
  1276
	     Rule.Thm ("sym_realpow_twoI",
walther@59871
  1277
                   ThmC.numerals_to_Free (@{thm realpow_twoI} RS @{thm sym})),
walther@60242
  1278
	     (*"r1 * r1 = r1 \<up> 2"*)
wneuper@59416
  1279
	     Rule.Thm ("sym_real_mult_2",
walther@59871
  1280
                   ThmC.numerals_to_Free (@{thm real_mult_2} RS @{thm sym})),
neuper@37950
  1281
	     (*"z1 + z1 = 2 * z1"*)
walther@59871
  1282
	     Rule.Thm ("real_mult_2_assoc",ThmC.numerals_to_Free @{thm real_mult_2_assoc}),
neuper@37950
  1283
	     (*"z1 + (z1 + k) = 2 * z1 + k"*)
neuper@37950
  1284
walther@59871
  1285
	     Rule.Thm ("real_num_collect",ThmC.numerals_to_Free @{thm real_num_collect}), 
neuper@37950
  1286
	     (*"[| l is_const; m is_const |]==>l * n + m * n = (l + m) * n"*)
walther@59871
  1287
	     Rule.Thm ("real_num_collect_assoc",ThmC.numerals_to_Free @{thm real_num_collect_assoc}),
neuper@37950
  1288
	     (*"[| l is_const; m is_const |] ==>  
neuper@37950
  1289
                                     l * n + (m * n + k) =  (l + m) * n + k"*)
walther@59871
  1290
	     Rule.Thm ("real_one_collect",ThmC.numerals_to_Free @{thm real_one_collect}),
neuper@37950
  1291
	     (*"m is_const ==> n + m * n = (1 + m) * n"*)
walther@59871
  1292
	     Rule.Thm ("real_one_collect_assoc",ThmC.numerals_to_Free @{thm real_one_collect_assoc}), 
neuper@37950
  1293
	     (*"m is_const ==> k + (n + m * n) = k + (1 + m) * n"*)
neuper@37950
  1294
walther@59871
  1295
	     Rule.Thm ("realpow_multI", ThmC.numerals_to_Free @{thm realpow_multI}),
walther@60242
  1296
	     (*"(r * s) \<up> n = r \<up> n * s \<up> n"*)
neuper@37950
  1297
walther@60278
  1298
	     Rule.Eval ("Groups.plus_class.plus", eval_binop "#add_"), 
walther@60278
  1299
	     Rule.Eval ("Groups.times_class.times", eval_binop "#mult_"),
walther@60278
  1300
	     Rule.Eval ("Transcendental.powr", eval_binop "#power_"),
neuper@37950
  1301
walther@59871
  1302
	     Rule.Thm ("mult_1_left",ThmC.numerals_to_Free @{thm mult_1_left}),(*"1 * z = z"*)
walther@59871
  1303
	     Rule.Thm ("mult_zero_left",ThmC.numerals_to_Free @{thm mult_zero_left}),(*"0 * z = 0"*)
walther@59871
  1304
	     Rule.Thm ("add_0_left",ThmC.numerals_to_Free @{thm add_0_left})(*0 + z = z*)
neuper@37950
  1305
wneuper@59416
  1306
	     (*Rule.Rls_ order_add_rls_*)
neuper@37950
  1307
	     ],
neuper@37950
  1308
walther@59878
  1309
    scr = Rule.Empty_Prog};      
wneuper@59472
  1310
\<close>
neuper@52125
  1311
wneuper@59523
  1312
subsection \<open>rule-sets with explicit program for intermediate steps\<close>
wneuper@59523
  1313
partial_function (tailrec) expand_binoms_2 :: "real \<Rightarrow> real"
wneuper@59523
  1314
  where
walther@59635
  1315
"expand_binoms_2 term = (
walther@59635
  1316
  Repeat (
walther@59637
  1317
    (Try (Repeat (Rewrite ''real_plus_binom_pow2''))) #>
walther@59637
  1318
    (Try (Repeat (Rewrite ''real_plus_binom_times''))) #>
walther@59637
  1319
    (Try (Repeat (Rewrite ''real_minus_binom_pow2''))) #>
walther@59637
  1320
    (Try (Repeat (Rewrite ''real_minus_binom_times''))) #>
walther@59637
  1321
    (Try (Repeat (Rewrite ''real_plus_minus_binom1''))) #>
walther@59637
  1322
    (Try (Repeat (Rewrite ''real_plus_minus_binom2''))) #>
walther@59635
  1323
   
walther@59637
  1324
    (Try (Repeat (Rewrite ''mult_1_left''))) #>
walther@59637
  1325
    (Try (Repeat (Rewrite ''mult_zero_left''))) #>
walther@59637
  1326
    (Try (Repeat (Rewrite ''add_0_left''))) #>
walther@59635
  1327
   
walther@59637
  1328
    (Try (Repeat (Calculate ''PLUS''))) #>
walther@59637
  1329
    (Try (Repeat (Calculate ''TIMES''))) #>
walther@59637
  1330
    (Try (Repeat (Calculate ''POWER''))) #>
walther@59635
  1331
   
walther@59637
  1332
    (Try (Repeat (Rewrite ''sym_realpow_twoI''))) #>
walther@59637
  1333
    (Try (Repeat (Rewrite ''realpow_plus_1''))) #>
walther@59637
  1334
    (Try (Repeat (Rewrite ''sym_real_mult_2''))) #>
walther@59637
  1335
    (Try (Repeat (Rewrite ''real_mult_2_assoc''))) #>
walther@59635
  1336
   
walther@59637
  1337
    (Try (Repeat (Rewrite ''real_num_collect''))) #>
walther@59637
  1338
    (Try (Repeat (Rewrite ''real_num_collect_assoc''))) #>
walther@59635
  1339
   
walther@59637
  1340
    (Try (Repeat (Rewrite ''real_one_collect''))) #>
walther@59637
  1341
    (Try (Repeat (Rewrite ''real_one_collect_assoc''))) #>
walther@59635
  1342
   
walther@59637
  1343
    (Try (Repeat (Calculate ''PLUS''))) #>
walther@59637
  1344
    (Try (Repeat (Calculate ''TIMES''))) #>
walther@59635
  1345
    (Try (Repeat (Calculate ''POWER''))))
walther@59635
  1346
  term)"
wneuper@59523
  1347
ML \<open>
wneuper@59523
  1348
val expand_binoms = 
walther@59851
  1349
  Rule_Def.Repeat{id = "expand_binoms", preconds = [], rew_ord = ("termlessI",termlessI),
walther@59851
  1350
      erls = Atools_erls, srls = Rule_Set.Empty,
walther@60278
  1351
      calc = [("PLUS"  , ("Groups.plus_class.plus", eval_binop "#add_")), 
walther@60278
  1352
	      ("TIMES" , ("Groups.times_class.times", eval_binop "#mult_")),
walther@60278
  1353
	      ("POWER", ("Transcendental.powr", eval_binop "#power_"))
wneuper@59523
  1354
	      ], errpatts = [],
wneuper@59523
  1355
      rules = [Rule.Thm ("real_plus_binom_pow2",
walther@59871
  1356
                     ThmC.numerals_to_Free @{thm real_plus_binom_pow2}),     
walther@60242
  1357
	       (*"(a + b) \<up> 2 = a \<up> 2 + 2 * a * b + b \<up> 2"*)
wneuper@59523
  1358
	       Rule.Thm ("real_plus_binom_times",
walther@59871
  1359
                     ThmC.numerals_to_Free @{thm real_plus_binom_times}),    
wneuper@59523
  1360
	      (*"(a + b)*(a + b) = ...*)
wneuper@59523
  1361
	       Rule.Thm ("real_minus_binom_pow2",
walther@59871
  1362
                     ThmC.numerals_to_Free @{thm real_minus_binom_pow2}),   
walther@60242
  1363
	       (*"(a - b) \<up> 2 = a \<up> 2 - 2 * a * b + b \<up> 2"*)
wneuper@59523
  1364
	       Rule.Thm ("real_minus_binom_times",
walther@59871
  1365
                     ThmC.numerals_to_Free @{thm real_minus_binom_times}),   
wneuper@59523
  1366
	       (*"(a - b)*(a - b) = ...*)
wneuper@59523
  1367
	       Rule.Thm ("real_plus_minus_binom1",
walther@59871
  1368
                     ThmC.numerals_to_Free @{thm real_plus_minus_binom1}),   
walther@60242
  1369
		(*"(a + b) * (a - b) = a \<up> 2 - b \<up> 2"*)
wneuper@59523
  1370
	       Rule.Thm ("real_plus_minus_binom2",
walther@59871
  1371
                     ThmC.numerals_to_Free @{thm real_plus_minus_binom2}),   
walther@60242
  1372
		(*"(a - b) * (a + b) = a \<up> 2 - b \<up> 2"*)
wneuper@59523
  1373
	       (*RL 020915*)
walther@59871
  1374
	       Rule.Thm ("real_pp_binom_times",ThmC.numerals_to_Free @{thm real_pp_binom_times}), 
wneuper@59523
  1375
		(*(a + b)*(c + d) = a*c + a*d + b*c + b*d*)
walther@59871
  1376
               Rule.Thm ("real_pm_binom_times",ThmC.numerals_to_Free @{thm real_pm_binom_times}), 
wneuper@59523
  1377
		(*(a + b)*(c - d) = a*c - a*d + b*c - b*d*)
walther@59871
  1378
               Rule.Thm ("real_mp_binom_times",ThmC.numerals_to_Free @{thm real_mp_binom_times}), 
wneuper@59523
  1379
		(*(a - b)*(c + d) = a*c + a*d - b*c - b*d*)
walther@59871
  1380
               Rule.Thm ("real_mm_binom_times",ThmC.numerals_to_Free @{thm real_mm_binom_times}), 
wneuper@59523
  1381
		(*(a - b)*(c - d) = a*c - a*d - b*c + b*d*)
walther@59871
  1382
	       Rule.Thm ("realpow_multI",ThmC.numerals_to_Free @{thm realpow_multI}),
walther@60260
  1383
		(*(a*b) \<up> n = a \<up> n * b \<up> n*)
walther@59871
  1384
	       Rule.Thm ("real_plus_binom_pow3",ThmC.numerals_to_Free @{thm real_plus_binom_pow3}),
walther@60260
  1385
	        (* (a + b) \<up> 3 = a \<up> 3 + 3*a \<up> 2*b + 3*a*b \<up> 2 + b \<up> 3 *)
wneuper@59523
  1386
	       Rule.Thm ("real_minus_binom_pow3",
walther@59871
  1387
                     ThmC.numerals_to_Free @{thm real_minus_binom_pow3}),
walther@60260
  1388
	        (* (a - b) \<up> 3 = a \<up> 3 - 3*a \<up> 2*b + 3*a*b \<up> 2 - b \<up> 3 *)
wneuper@59523
  1389
wneuper@59523
  1390
walther@59871
  1391
              (*Rule.Thm ("distrib_right" ,ThmC.numerals_to_Free @{thm distrib_right}),	
wneuper@59523
  1392
		(*"(z1.0 + z2.0) * w = z1.0 * w + z2.0 * w"*)
walther@59871
  1393
	       Rule.Thm ("distrib_left",ThmC.numerals_to_Free @{thm distrib_left}),	
wneuper@59523
  1394
	       (*"w * (z1.0 + z2.0) = w * z1.0 + w * z2.0"*)
walther@59871
  1395
	       Rule.Thm ("left_diff_distrib" ,ThmC.numerals_to_Free @{thm left_diff_distrib}),	
wneuper@59523
  1396
	       (*"(z1.0 - z2.0) * w = z1.0 * w - z2.0 * w"*)
walther@59871
  1397
	       Rule.Thm ("right_diff_distrib",ThmC.numerals_to_Free @{thm right_diff_distrib}),	
wneuper@59523
  1398
	       (*"w * (z1.0 - z2.0) = w * z1.0 - w * z2.0"*)
wneuper@59523
  1399
	      *)
walther@59871
  1400
	       Rule.Thm ("mult_1_left",ThmC.numerals_to_Free @{thm mult_1_left}),
wneuper@59523
  1401
               (*"1 * z = z"*)
walther@59871
  1402
	       Rule.Thm ("mult_zero_left",ThmC.numerals_to_Free @{thm mult_zero_left}),
wneuper@59523
  1403
               (*"0 * z = 0"*)
walther@59871
  1404
	       Rule.Thm ("add_0_left",ThmC.numerals_to_Free @{thm add_0_left}),(*"0 + z = z"*)
wneuper@59523
  1405
walther@60278
  1406
	       Rule.Eval ("Groups.plus_class.plus", eval_binop "#add_"), 
walther@60278
  1407
	       Rule.Eval ("Groups.times_class.times", eval_binop "#mult_"),
walther@60278
  1408
	       Rule.Eval ("Transcendental.powr", eval_binop "#power_"),
walther@59877
  1409
              (*Rule.Thm ("mult.commute",ThmC.numerals_to_Free @{thm mult.commute}),
wneuper@59523
  1410
		(*AC-rewriting*)
wneuper@59523
  1411
	       Rule.Thm ("real_mult_left_commute",
walther@59871
  1412
                     ThmC.numerals_to_Free @{thm real_mult_left_commute}),
walther@59877
  1413
	       Rule.Thm ("mult.assoc",ThmC.numerals_to_Free @{thm mult.assoc}),
walther@59877
  1414
	       Rule.Thm ("add.commute",ThmC.numerals_to_Free @{thm add.commute}),
walther@59877
  1415
	       Rule.Thm ("add.left_commute",ThmC.numerals_to_Free @{thm add.left_commute}),
walther@59877
  1416
	       Rule.Thm ("add.assoc",ThmC.numerals_to_Free @{thm add.assoc}),
wneuper@59523
  1417
	      *)
wneuper@59523
  1418
	       Rule.Thm ("sym_realpow_twoI",
walther@59871
  1419
                     ThmC.numerals_to_Free (@{thm realpow_twoI} RS @{thm sym})),
walther@60242
  1420
	       (*"r1 * r1 = r1 \<up> 2"*)
walther@59871
  1421
	       Rule.Thm ("realpow_plus_1",ThmC.numerals_to_Free @{thm realpow_plus_1}),			
walther@60242
  1422
	       (*"r * r \<up> n = r \<up> (n + 1)"*)
wneuper@59523
  1423
	       (*Rule.Thm ("sym_real_mult_2",
walther@59871
  1424
                       ThmC.numerals_to_Free (@{thm real_mult_2} RS @{thm sym})),		
wneuper@59523
  1425
	       (*"z1 + z1 = 2 * z1"*)*)
walther@59871
  1426
	       Rule.Thm ("real_mult_2_assoc",ThmC.numerals_to_Free @{thm real_mult_2_assoc}),		
wneuper@59523
  1427
	       (*"z1 + (z1 + k) = 2 * z1 + k"*)
wneuper@59523
  1428
walther@59871
  1429
	       Rule.Thm ("real_num_collect",ThmC.numerals_to_Free @{thm real_num_collect}), 
wneuper@59523
  1430
	       (*"[| l is_const; m is_const |] ==>l * n + m * n = (l + m) * n"*)
wneuper@59523
  1431
	       Rule.Thm ("real_num_collect_assoc",
walther@59871
  1432
                     ThmC.numerals_to_Free @{thm real_num_collect_assoc}),	
wneuper@59523
  1433
	       (*"[| l is_const; m is_const |] ==>  
wneuper@59523
  1434
                                       l * n + (m * n + k) =  (l + m) * n + k"*)
walther@59871
  1435
	       Rule.Thm ("real_one_collect",ThmC.numerals_to_Free @{thm real_one_collect}),
wneuper@59523
  1436
	       (*"m is_const ==> n + m * n = (1 + m) * n"*)
wneuper@59523
  1437
	       Rule.Thm ("real_one_collect_assoc",
walther@59871
  1438
                     ThmC.numerals_to_Free @{thm real_one_collect_assoc}), 
wneuper@59523
  1439
	       (*"m is_const ==> k + (n + m * n) = k + (1 + m) * n"*)
wneuper@59523
  1440
walther@60278
  1441
	       Rule.Eval ("Groups.plus_class.plus", eval_binop "#add_"), 
walther@60278
  1442
	       Rule.Eval ("Groups.times_class.times", eval_binop "#mult_"),
walther@60278
  1443
	       Rule.Eval ("Transcendental.powr", eval_binop "#power_")
wneuper@59523
  1444
	       ],
walther@59618
  1445
      scr = Rule.Prog (Program.prep_program @{thm expand_binoms_2.simps})
wneuper@59523
  1446
      };      
wneuper@59523
  1447
\<close>
wneuper@59523
  1448
walther@59887
  1449
subsection \<open>add to Know_Store\<close>
wneuper@59523
  1450
subsubsection \<open>rule-sets\<close>
walther@59618
  1451
ML \<open>val prep_rls' = Auto_Prog.prep_rls @{theory}\<close>
s1210629013@55444
  1452
wneuper@59472
  1453
setup \<open>KEStore_Elems.add_rlss 
s1210629013@55444
  1454
  [("norm_Poly", (Context.theory_name @{theory}, prep_rls' norm_Poly)), 
s1210629013@55444
  1455
  ("Poly_erls", (Context.theory_name @{theory}, prep_rls' Poly_erls)),(*FIXXXME:del with rls.rls'*) 
s1210629013@55444
  1456
  ("expand", (Context.theory_name @{theory}, prep_rls' expand)), 
s1210629013@55444
  1457
  ("expand_poly", (Context.theory_name @{theory}, prep_rls' expand_poly)), 
s1210629013@55444
  1458
  ("simplify_power", (Context.theory_name @{theory}, prep_rls' simplify_power)),
neuper@52125
  1459
s1210629013@55444
  1460
  ("order_add_mult", (Context.theory_name @{theory}, prep_rls' order_add_mult)), 
s1210629013@55444
  1461
  ("collect_numerals", (Context.theory_name @{theory}, prep_rls' collect_numerals)), 
s1210629013@55444
  1462
  ("collect_numerals_", (Context.theory_name @{theory}, prep_rls' collect_numerals_)), 
s1210629013@55444
  1463
  ("reduce_012", (Context.theory_name @{theory}, prep_rls' reduce_012)), 
s1210629013@55444
  1464
  ("discard_parentheses", (Context.theory_name @{theory}, prep_rls' discard_parentheses)),
neuper@52125
  1465
 
s1210629013@55444
  1466
  ("make_polynomial", (Context.theory_name @{theory}, prep_rls' make_polynomial)), 
s1210629013@55444
  1467
  ("expand_binoms", (Context.theory_name @{theory}, prep_rls' expand_binoms)), 
s1210629013@55444
  1468
  ("rev_rew_p", (Context.theory_name @{theory}, prep_rls' rev_rew_p)), 
s1210629013@55444
  1469
  ("discard_minus", (Context.theory_name @{theory}, prep_rls' discard_minus)), 
s1210629013@55444
  1470
  ("expand_poly_", (Context.theory_name @{theory}, prep_rls' expand_poly_)),
neuper@52125
  1471
 
s1210629013@55444
  1472
  ("expand_poly_rat_", (Context.theory_name @{theory}, prep_rls' expand_poly_rat_)), 
s1210629013@55444
  1473
  ("simplify_power_", (Context.theory_name @{theory}, prep_rls' simplify_power_)), 
s1210629013@55444
  1474
  ("calc_add_mult_pow_", (Context.theory_name @{theory}, prep_rls' calc_add_mult_pow_)), 
s1210629013@55444
  1475
  ("reduce_012_mult_", (Context.theory_name @{theory}, prep_rls' reduce_012_mult_)), 
s1210629013@55444
  1476
  ("reduce_012_", (Context.theory_name @{theory}, prep_rls' reduce_012_)),
neuper@52125
  1477
 
s1210629013@55444
  1478
  ("discard_parentheses1", (Context.theory_name @{theory}, prep_rls' discard_parentheses1)), 
s1210629013@55444
  1479
  ("order_mult_rls_", (Context.theory_name @{theory}, prep_rls' order_mult_rls_)), 
s1210629013@55444
  1480
  ("order_add_rls_", (Context.theory_name @{theory}, prep_rls' order_add_rls_)), 
neuper@52125
  1481
  ("make_rat_poly_with_parentheses",
wneuper@59472
  1482
    (Context.theory_name @{theory}, prep_rls' make_rat_poly_with_parentheses))]\<close>
wneuper@59523
  1483
wneuper@59526
  1484
subsection \<open>problems\<close>
wneuper@59472
  1485
setup \<open>KEStore_Elems.add_pbts
walther@59973
  1486
  [(Problem.prep_input thy "pbl_simp_poly" [] Problem.id_empty
walther@59997
  1487
      (["polynomial", "simplification"],
s1210629013@55339
  1488
        [("#Given" ,["Term t_t"]),
s1210629013@55339
  1489
          ("#Where" ,["t_t is_polyexp"]),
s1210629013@55339
  1490
          ("#Find"  ,["normalform n_n"])],
walther@59852
  1491
        Rule_Set.append_rules "empty" Rule_Set.empty [(*for preds in where_*)
walther@60278
  1492
			  Rule.Eval ("Poly.is_polyexp", eval_is_polyexp "")], 
s1210629013@55339
  1493
        SOME "Simplify t_t", 
walther@59997
  1494
        [["simplification", "for_polynomials"]]))]\<close>
wneuper@59429
  1495
wneuper@59526
  1496
subsection \<open>methods\<close>
wneuper@59545
  1497
wneuper@59429
  1498
partial_function (tailrec) simplify :: "real \<Rightarrow> real"
wneuper@59429
  1499
  where
walther@59635
  1500
"simplify term = ((Rewrite_Set ''norm_Poly'') term)"
wneuper@59472
  1501
setup \<open>KEStore_Elems.add_mets
walther@60154
  1502
    [MethodC.prep_input thy "met_simp_poly" [] MethodC.id_empty
walther@59997
  1503
	    (["simplification", "for_polynomials"],
s1210629013@55373
  1504
	      [("#Given" ,["Term t_t"]),
s1210629013@55373
  1505
	        ("#Where" ,["t_t is_polyexp"]),
s1210629013@55373
  1506
	        ("#Find"  ,["normalform n_n"])],
walther@59852
  1507
	      {rew_ord'="tless_true", rls' = Rule_Set.empty, calc = [], srls = Rule_Set.empty, 
walther@59852
  1508
	        prls = Rule_Set.append_rules "simplification_for_polynomials_prls" Rule_Set.empty 
s1210629013@55373
  1509
				    [(*for preds in where_*)
walther@60278
  1510
				      Rule.Eval ("Poly.is_polyexp", eval_is_polyexp"")],
walther@59852
  1511
				  crls = Rule_Set.empty, errpats = [], nrls = norm_Poly},
wneuper@59552
  1512
        @{thm simplify.simps})]
wneuper@59472
  1513
\<close>
wneuper@59472
  1514
ML \<open>
wneuper@59472
  1515
\<close> ML \<open>
wneuper@59472
  1516
\<close> 
neuper@37906
  1517
end