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