src/Tools/isac/Knowledge/Root.thy
author wneuper <Walther.Neuper@jku.at>
Wed, 11 Jan 2023 11:38:01 +0100
changeset 60650 06ec8abfd3bc
parent 60594 439f7f3867ec
child 60671 8998cb4818dd
permissions -rw-r--r--
eliminate use of Thy_Info 12: TermC partially
neuper@37906
     1
(* theory collecting all knowledge for Root
neuper@37906
     2
   created by: 
neuper@37906
     3
         date: 
neuper@37906
     4
   changed by: rlang
neuper@37906
     5
   last change by: rlang
neuper@37906
     6
             date: 02.10.21
neuper@37906
     7
*)
neuper@37906
     8
neuper@37982
     9
theory Root imports Poly begin
neuper@37906
    10
neuper@37906
    11
consts
neuper@37906
    12
neuper@37982
    13
  (*sqrt   :: "real => real"         Isabelle "NthRoot.sqrt"*)
neuper@37950
    14
  nroot  :: "[real, real] => real"
neuper@37906
    15
neuper@52148
    16
axiomatization where (*.not contained in Isabelle2002,
neuper@37906
    17
         stated as axioms, TODO: prove as theorems;
neuper@37906
    18
         theorem-IDs 'xxxI' with ^^^ instead of ^ in 'xxx' in Isabelle2002.*)
neuper@37906
    19
neuper@37982
    20
  root_plus_minus:         "0 <= b ==> 
walther@60242
    21
			   (a \<up> 2 = b) = ((a = sqrt b) | (a = (-1)*sqrt b))" and
walther@60242
    22
  root_false:		  "b < 0 ==> (a \<up> 2 = b) = False" and
neuper@37906
    23
neuper@37906
    24
 (* for expand_rootbinom *)
neuper@52148
    25
  real_pp_binom_times:     "(a + b)*(c + d) = a*c + a*d + b*c + b*d" and
neuper@52148
    26
  real_pm_binom_times:     "(a + b)*(c - d) = a*c - a*d + b*c - b*d" and
neuper@52148
    27
  real_mp_binom_times:     "(a - b)*(c + d) = a*c + a*d - b*c - b*d" and
neuper@52148
    28
  real_mm_binom_times:     "(a - b)*(c - d) = a*c - a*d - b*c + b*d" and
walther@60242
    29
  real_plus_binom_pow3:    "(a + b) \<up> 3 = a \<up> 3 + 3*a \<up> 2*b + 3*a*b \<up> 2 + b \<up> 3" and
walther@60242
    30
  real_minus_binom_pow3:   "(a - b) \<up> 3 = a \<up> 3 - 3*a \<up> 2*b + 3*a*b \<up> 2 - b \<up> 3" and
walther@60242
    31
  realpow_mul:             "(a*b) \<up> n = a \<up> n * b \<up> n" and
neuper@37906
    32
neuper@52148
    33
  real_diff_minus:         "a - b = a + (-1) * b" and
walther@60242
    34
  real_plus_binom_times:   "(a + b)*(a + b) = a \<up> 2 + 2*a*b + b \<up> 2" and
walther@60242
    35
  real_minus_binom_times:  "(a - b)*(a - b) = a \<up> 2 - 2*a*b + b \<up> 2" and
walther@60242
    36
  real_plus_binom_pow2:    "(a + b) \<up> 2 = a \<up> 2 + 2*a*b + b \<up> 2" and
walther@60242
    37
  real_minus_binom_pow2:   "(a - b) \<up> 2 = a \<up> 2 - 2*a*b + b \<up> 2" and
walther@60242
    38
  real_plus_minus_binom1:  "(a + b)*(a - b) = a \<up> 2 - b \<up> 2" and
walther@60242
    39
  real_plus_minus_binom2:  "(a - b)*(a + b) = a \<up> 2 - b \<up> 2" and
neuper@37906
    40
walther@60242
    41
  real_root_positive:      "0 <= a ==> (x  \<up>  2 = a) = (x = sqrt a)" and
walther@60242
    42
  real_root_negative:      "a <  0 ==> (x  \<up>  2 = a) = False"
neuper@37906
    43
wneuper@59472
    44
ML \<open>
neuper@37950
    45
(*-------------------------functions---------------------*)
neuper@37950
    46
(*evaluation square-root over the integers*)
Walther@60504
    47
fun eval_sqrt (_ : string) (_ : string) t (_: Proof.context) = 
wenzelm@60381
    48
  (case t of
wenzelm@60381
    49
    Const (op0, _) $ num =>
wenzelm@60381
    50
      (case try HOLogic.dest_number num of
wenzelm@60381
    51
        SOME (T, ni)=>
wenzelm@60381
    52
          if ni < 0 then NONE
wenzelm@60381
    53
          else
wenzelm@60381
    54
            let val fact = Eval.squfact ni;
wenzelm@60381
    55
            in
walther@60317
    56
              if fact * fact = ni 
wenzelm@60381
    57
              then
walther@60317
    58
                SOME ("#sqrt #" ^ string_of_int ni ^ " = #"
wenzelm@60381
    59
                    ^ string_of_int (if ni = 0 then 0 else ni div fact),
wenzelm@60381
    60
                  HOLogic.Trueprop $ TermC.mk_equality (t, TermC.term_of_num T fact))
wenzelm@60381
    61
              else if fact = 1 then NONE
wenzelm@60381
    62
              else
walther@60317
    63
                SOME ("#sqrt #" ^ string_of_int ni ^ " = sqrt (#"
wenzelm@60381
    64
                    ^ string_of_int fact ^ " * #" ^ string_of_int fact ^ " * #"
wenzelm@60381
    65
                    ^ string_of_int (ni div (fact * fact)) ^ ")",
wenzelm@60381
    66
                  HOLogic.Trueprop $ TermC.mk_equality 
walther@60317
    67
                    (t, (TermC.mk_factroot op0 T fact (ni div (fact*fact)))))
wenzelm@60381
    68
            end
wenzelm@60381
    69
        | NONE => NONE)
wenzelm@60381
    70
  | _ => NONE);
wenzelm@60381
    71
Walther@60567
    72
(*val (thmid, op_, t as Const(op0,t0) $ arg) = ("", "", TermC.parse_test @{context} "sqrt 0");
neuper@37950
    73
> eval_sqrt thmid op_ t thy;
neuper@37950
    74
> val Free (n1,t1) = arg; 
neuper@37950
    75
> val SOME ni = int_of_str n1;
neuper@37950
    76
*)
wneuper@59472
    77
\<close>
wenzelm@60381
    78
wenzelm@60313
    79
calculation SQRT (sqrt) = \<open>eval_sqrt "#sqrt_"\<close>
wenzelm@60313
    80
    (*different types for 'sqrt 4' --- 'Calculate SQRT'*)
wneuper@59472
    81
ML \<open>
neuper@37950
    82
local (* Vers. 7.10.99.A *)
neuper@37950
    83
neuper@37950
    84
open Term;  (* for type order = EQUAL | LESS | GREATER *)
neuper@37950
    85
neuper@37950
    86
fun pr_ord EQUAL = "EQUAL"
neuper@37950
    87
  | pr_ord LESS  = "LESS"
neuper@37950
    88
  | pr_ord GREATER = "GREATER";
neuper@37950
    89
neuper@37950
    90
fun dest_hd' (Const (a, T)) =                          (* ~ term.ML *)
wenzelm@60309
    91
  (case a of \<^const_name>\<open>sqrt\<close>  => ((("|||", 0), T), 0)      (*WN greatest *)
neuper@37950
    92
	   | _ => (((a, 0), T), 0))
walther@60317
    93
  | dest_hd' (Free (a, T)) = (((a, 0), T), 1)(*TODOO handle this as numeral, too? see EqSystem.thy*)
neuper@37950
    94
  | dest_hd' (Var v) = (v, 2)
neuper@37950
    95
  | dest_hd' (Bound i) = ((("", i), dummyT), 3)
walther@60269
    96
  | dest_hd' (Abs (_, T, _)) = ((("", 0), T), 4)
walther@60269
    97
  | dest_hd' _ = raise ERROR "dest_hd': uncovered case in fun.def.";
neuper@37950
    98
fun size_of_term' (Const(str,_) $ t) =
wenzelm@60309
    99
    (case str of \<^const_name>\<open>sqrt\<close>  => (1000 + size_of_term' t)
neuper@37950
   100
               | _ => 1 + size_of_term' t)
neuper@37950
   101
  | size_of_term' (Abs (_,_,body)) = 1 + size_of_term' body
neuper@37950
   102
  | size_of_term' (f $ t) = size_of_term' f  +  size_of_term' t
neuper@37950
   103
  | size_of_term' _ = 1;
Walther@60594
   104
fun term_ord' pr ctxt (Abs (_, T, t), Abs(_, U, u)) =       (* ~ term.ML *)
Walther@60594
   105
    (case term_ord' pr ctxt (t, u) of EQUAL => Term_Ord.typ_ord (T, U) 
neuper@38053
   106
                                   | ord => ord)
Walther@60594
   107
  | term_ord' pr ctxt (t, u) =
Walther@60594
   108
    (if pr 
Walther@60594
   109
     then 
Walther@60594
   110
  	   let
Walther@60594
   111
  	     val (f, ts) = strip_comb t and (g, us) = strip_comb u;
Walther@60594
   112
  	     val _ = tracing ("t= f@ts= \"" ^ UnparseC.term_in_ctxt ctxt f ^"\" @ \"[" ^
Walther@60594
   113
  	                      commas (map (UnparseC.term_in_ctxt ctxt) ts) ^ "]\"");
Walther@60594
   114
  	     val _ = tracing ("u= g@us= \"" ^ UnparseC.term g ^"\" @ \"[" ^
Walther@60594
   115
  	                      commas (map (UnparseC.term_in_ctxt ctxt) us) ^ "]\"");
Walther@60594
   116
  	     val _ = tracing ("size_of_term(t,u)= (" ^
Walther@60594
   117
  	                      string_of_int(size_of_term' t) ^", " ^
Walther@60594
   118
  	                      string_of_int(size_of_term' u) ^")");
Walther@60594
   119
  	     val _ = tracing ("hd_ord(f,g)      = " ^ (pr_ord o hd_ord) (f,g));
Walther@60594
   120
  	     val _ = tracing ("terms_ord(ts,us) = " ^ (pr_ord o terms_ord str false ctxt) (ts,us));
Walther@60594
   121
  	     val _=tracing("-------");
Walther@60594
   122
  	   in () end
neuper@38053
   123
     else ();
Walther@60594
   124
    case int_ord (size_of_term' t, size_of_term' u) of
Walther@60594
   125
	    EQUAL =>
Walther@60594
   126
	      let val (f, ts) = strip_comb t and (g, us) = strip_comb u 
Walther@60594
   127
        in (case hd_ord (f, g) of
Walther@60594
   128
              EQUAL => (terms_ord str pr ctxt) (ts, us) 
Walther@60594
   129
	          | ord => ord)
Walther@60594
   130
	      end
Walther@60594
   131
    | ord => ord)
neuper@37950
   132
and hd_ord (f, g) =                                        (* ~ term.ML *)
Walther@60594
   133
  prod_ord (prod_ord Term_Ord.indexname_ord Term_Ord.typ_ord) int_ord  (dest_hd' f, dest_hd' g)
Walther@60594
   134
and terms_ord _(*str*) pr ctxt (ts, us) = list_ord (term_ord' pr ctxt) (ts, us);
neuper@37950
   135
neuper@37950
   136
in
neuper@37950
   137
(* associates a+(b+c) => (a+b)+c = a+b+c ... avoiding parentheses 
neuper@37950
   138
  by (1) size_of_term: less(!) to right, size_of 'sqrt (...)' = 1 
neuper@37950
   139
     (2) hd_ord: greater to right, 'sqrt' < numerals < variables
neuper@37950
   140
     (3) terms_ord: recurs. on args, greater to right
neuper@37950
   141
*)
neuper@37950
   142
neuper@37950
   143
(*args
neuper@37950
   144
   pr: print trace, WN0509 'sqrt_right true' not used anymore
Walther@60594
   145
   ctxt:
neuper@37950
   146
   subst: no bound variables, only Root.sqrt
neuper@37950
   147
   tu: the terms to compare (t1, t2) ... *)
Walther@60594
   148
fun sqrt_right (pr:bool) ctxt (_: subst) (ts, us) = 
Walther@60594
   149
    (term_ord' pr ctxt (TermC.numerals_to_Free ts, TermC.numerals_to_Free us) = LESS );
neuper@37950
   150
end;
Walther@60506
   151
\<close> ML \<open>
Walther@60594
   152
\<close> ML \<open>
Walther@60594
   153
sqrt_right: bool -> Proof.context -> Subst.T -> term * term -> bool
Walther@60506
   154
\<close> 
Walther@60588
   155
setup \<open>Know_Store.add_rew_ords [
Walther@60506
   156
  ("termlessI", termlessI),
Walther@60594
   157
  ("sqrt_right", sqrt_right false)]\<close>
neuper@37950
   158
Walther@60506
   159
ML \<open>
Walther@60506
   160
(*------------------------- rules -------------------------*)
neuper@37950
   161
val Root_crls = 
walther@60358
   162
  Rule_Set.append_rules "Root_crls" Atools_erls [
walther@60358
   163
    \<^rule_thm>\<open>real_unari_minus\<close>,
walther@60358
   164
    \<^rule_eval>\<open>sqrt\<close> (eval_sqrt "#sqrt_"),
walther@60358
   165
    \<^rule_eval>\<open>divide\<close> (Prog_Expr.eval_cancel "#divide_e"),
Walther@60515
   166
    \<^rule_eval>\<open>realpow\<close> (**)(Calc_Binop.numeric "#power_"),
Walther@60515
   167
    \<^rule_eval>\<open>plus\<close> (**)(Calc_Binop.numeric "#add_"), 
Walther@60515
   168
    \<^rule_eval>\<open>minus\<close> (**)(Calc_Binop.numeric "#sub_"),
Walther@60515
   169
    \<^rule_eval>\<open>times\<close> (**)(Calc_Binop.numeric "#mult_"),
walther@60358
   170
    \<^rule_eval>\<open>HOL.eq\<close> (Prog_Expr.eval_equal "#equal_")];
neuper@37950
   171
neuper@37950
   172
val Root_erls = 
walther@60358
   173
  Rule_Set.append_rules "Root_erls" Atools_erls [
walther@60358
   174
    \<^rule_thm>\<open>real_unari_minus\<close>,
walther@60358
   175
    \<^rule_eval>\<open>sqrt\<close> (eval_sqrt "#sqrt_"),
walther@60358
   176
    \<^rule_eval>\<open>Rings.divide_class.divide\<close> (Prog_Expr.eval_cancel "#divide_e"),
Walther@60515
   177
    \<^rule_eval>\<open>realpow\<close> (**)(Calc_Binop.numeric "#power_"),
Walther@60515
   178
    \<^rule_eval>\<open>Groups.plus_class.plus\<close> (**)(Calc_Binop.numeric "#add_"), 
Walther@60515
   179
    \<^rule_eval>\<open>Groups.minus_class.minus\<close> (**)(Calc_Binop.numeric "#sub_"),
Walther@60515
   180
    \<^rule_eval>\<open>Groups.times_class.times\<close> (**)(Calc_Binop.numeric "#mult_"),
walther@60358
   181
    \<^rule_eval>\<open>HOL.eq\<close> (Prog_Expr.eval_equal "#equal_")];
walther@60384
   182
val Root_erls = 
walther@60384
   183
  Rule_Set.append_erls_rules "Root_erls_erls" Root_erls [
walther@60384
   184
    \<^rule_thm>\<open>not_false\<close>,
walther@60384
   185
    \<^rule_thm>\<open>not_true\<close>,
walther@60384
   186
    \<^rule_eval>\<open>is_num\<close> (Prog_Expr.eval_is_num "#is_num_")];
wneuper@59472
   187
\<close>
wenzelm@60289
   188
rule_set_knowledge Root_erls = Root_erls
wneuper@59472
   189
ML \<open>
neuper@37950
   190
s1210629013@55444
   191
val make_rooteq = prep_rls'(
walther@59851
   192
  Rule_Def.Repeat{id = "make_rooteq", preconds = []:term list, 
Walther@60594
   193
    rew_ord = ("sqrt_right", sqrt_right false),
Walther@60586
   194
    asm_rls = Atools_erls, prog_rls = Rule_Set.Empty,
walther@60358
   195
    calc = [], errpatts = [],
walther@60358
   196
    rules = [
walther@60358
   197
      \<^rule_thm>\<open>real_diff_minus\<close>,	 (*"a - b = a + (-1) * b"*)
neuper@37950
   198
walther@60358
   199
      \<^rule_thm>\<open>distrib_right\<close>,	 (*"(z1.0 + z2.0) * w = z1.0 * w + z2.0 * w"*)
walther@60358
   200
      \<^rule_thm>\<open>distrib_left\<close>, (*"w * (z1.0 + z2.0) = w * z1.0 + w * z2.0"*)
walther@60358
   201
      \<^rule_thm>\<open>left_diff_distrib\<close>, (*"(z1.0 - z2.0) * w = z1.0 * w - z2.0 * w"*)
walther@60358
   202
      \<^rule_thm>\<open>right_diff_distrib\<close>, (*"w * (z1.0 - z2.0) = w * z1.0 - w * z2.0"*)
neuper@37950
   203
walther@60358
   204
      \<^rule_thm>\<open>mult_1_left\<close>, (*"1 * z = z"*)
walther@60358
   205
      \<^rule_thm>\<open>mult_zero_left\<close>, (*"0 * z = 0"*)
walther@60358
   206
      \<^rule_thm>\<open>add_0_left\<close>, (*"0 + z = z"*)
neuper@37950
   207
walther@60358
   208
       (*AC-rewriting*)
walther@60358
   209
      \<^rule_thm>\<open>mult.commute\<close>,
walther@60358
   210
      \<^rule_thm>\<open>real_mult_left_commute\<close>,
walther@60358
   211
      \<^rule_thm>\<open>mult.assoc\<close>,
walther@60358
   212
      \<^rule_thm>\<open>add.commute\<close>,
walther@60358
   213
      \<^rule_thm>\<open>add.left_commute\<close>,
walther@60358
   214
      \<^rule_thm>\<open>add.assoc\<close>,
neuper@37950
   215
walther@60358
   216
      \<^rule_thm_sym>\<open>realpow_twoI\<close>, (*"r1 * r1 = r1  \<up>  2"*)
walther@60358
   217
      \<^rule_thm>\<open>realpow_plus_1\<close>, (*"r * r  \<up>  n = r  \<up>  (n + 1)"*)
walther@60358
   218
      \<^rule_thm_sym>\<open>real_mult_2\<close>, (*"z1 + z1 = 2 * z1"*)
walther@60358
   219
      \<^rule_thm>\<open>real_mult_2_assoc\<close>,	(*"z1 + (z1 + k) = 2 * z1 + k"*)
neuper@37950
   220
walther@60385
   221
      \<^rule_thm>\<open>real_num_collect\<close>, (*"[| l is_num; m is_num |]==> l * n + m * n = (l + m) * n"*)
walther@60385
   222
      \<^rule_thm>\<open>real_num_collect_assoc\<close>, (*"[| l is_num; m is_num |] ==> l * n + (m * n + k) =  (l + m) * n + k"*)
walther@60385
   223
      \<^rule_thm>\<open>real_one_collect\<close>, (*"m is_num ==> n + m * n = (1 + m) * n"*)
walther@60385
   224
      \<^rule_thm>\<open>real_one_collect_assoc\<close>, (*"m is_num ==> k + (n + m * n) = k + (1 + m) * n"*)
walther@60358
   225
Walther@60515
   226
      \<^rule_eval>\<open>plus\<close> (**)(Calc_Binop.numeric "#add_"),
Walther@60515
   227
      \<^rule_eval>\<open>times\<close> (**)(Calc_Binop.numeric "#mult_"),
Walther@60515
   228
      \<^rule_eval>\<open>realpow\<close> (**)(Calc_Binop.numeric "#power_")],
Walther@60586
   229
    program = Rule.Empty_Prog});      
wneuper@59472
   230
\<close>
wenzelm@60289
   231
rule_set_knowledge make_rooteq = make_rooteq
wneuper@59472
   232
ML \<open>
neuper@37950
   233
walther@59618
   234
val prep_rls' = Auto_Prog.prep_rls @{theory};
s1210629013@55444
   235
s1210629013@55444
   236
val expand_rootbinoms = prep_rls'(
walther@60358
   237
  Rule_Def.Repeat {
walther@60358
   238
    id = "expand_rootbinoms", preconds = [], rew_ord = ("termlessI",termlessI),
Walther@60586
   239
    asm_rls = Atools_erls, prog_rls = Rule_Set.Empty, calc = [], errpatts = [],
walther@60358
   240
    rules = [
walther@60358
   241
       \<^rule_thm>\<open>real_plus_binom_pow2\<close>, (*"(a + b)  \<up>  2 = a  \<up>  2 + 2 * a * b + b  \<up>  2"*)
walther@60358
   242
       \<^rule_thm>\<open>real_plus_binom_times\<close>, (*"(a + b)*(a + b) = ...*)
walther@60358
   243
       \<^rule_thm>\<open>real_minus_binom_pow2\<close>, (*"(a - b)  \<up>  2 = a  \<up>  2 - 2 * a * b + b  \<up>  2"*)
walther@60358
   244
       \<^rule_thm>\<open>real_minus_binom_times\<close>, (*"(a - b)*(a - b) = ...*)
walther@60358
   245
       \<^rule_thm>\<open>real_plus_minus_binom1\<close>, (*"(a + b) * (a - b) = a  \<up>  2 - b  \<up>  2"*)
walther@60358
   246
       \<^rule_thm>\<open>real_plus_minus_binom2\<close>,  (*"(a - b) * (a + b) = a  \<up>  2 - b  \<up>  2"*)
walther@60358
   247
       (*RL 020915*)
walther@60358
   248
       \<^rule_thm>\<open>real_pp_binom_times\<close>, (*(a + b)*(c + d) = a*c + a*d + b*c + b*d*)
walther@60358
   249
       \<^rule_thm>\<open>real_pm_binom_times\<close>, (*(a + b)*(c - d) = a*c - a*d + b*c - b*d*)
walther@60358
   250
       \<^rule_thm>\<open>real_mp_binom_times\<close>, (*(a - b)*(c p d) = a*c + a*d - b*c - b*d*)
walther@60358
   251
       \<^rule_thm>\<open>real_mm_binom_times\<close>, (*(a - b)*(c p d) = a*c - a*d - b*c + b*d*)
walther@60358
   252
       \<^rule_thm>\<open>realpow_mul\<close>,  (*(a*b) \<up> n = a \<up> n * b \<up> n*)
walther@60358
   253
  
walther@60358
   254
       \<^rule_thm>\<open>mult_1_left\<close>, (*"1 * z = z"*)
walther@60358
   255
       \<^rule_thm>\<open>mult_zero_left\<close>, (*"0 * z = 0"*)
walther@60358
   256
       \<^rule_thm>\<open>add_0_left\<close>, (*"0 + z = z"*)
walther@60358
   257
  
Walther@60515
   258
       \<^rule_eval>\<open>plus\<close> (**)(Calc_Binop.numeric "#add_"), 
Walther@60515
   259
       \<^rule_eval>\<open>minus\<close> (**)(Calc_Binop.numeric "#sub_"), 
Walther@60515
   260
       \<^rule_eval>\<open>times\<close> (**)(Calc_Binop.numeric "#mult_"),
walther@60358
   261
       \<^rule_eval>\<open>divide\<close> (Prog_Expr.eval_cancel "#divide_e"),
walther@60358
   262
       \<^rule_eval>\<open>sqrt\<close> (eval_sqrt "#sqrt_"),
Walther@60515
   263
       \<^rule_eval>\<open>realpow\<close> (**)(Calc_Binop.numeric "#power_"),
walther@60358
   264
  
walther@60358
   265
       \<^rule_thm_sym>\<open>realpow_twoI\<close>, (*"r1 * r1 = r1  \<up>  2"*)
walther@60358
   266
       \<^rule_thm>\<open>realpow_plus_1\<close>, (*"r * r  \<up>  n = r  \<up>  (n + 1)"*)
walther@60358
   267
       \<^rule_thm>\<open>real_mult_2_assoc\<close>,	 (*"z1 + (z1 + k) = 2 * z1 + k"*)
walther@60358
   268
  
walther@60385
   269
       \<^rule_thm>\<open>real_num_collect\<close>, (*"[| l is_num; m is_num |] ==>l * n + m * n = (l + m) * n"*)
walther@60385
   270
       \<^rule_thm>\<open>real_num_collect_assoc\<close>, (*"[| l is_num; m is_num |] ==> l * n + (m * n + k) =  (l + m) * n + k"*)
walther@60385
   271
       \<^rule_thm>\<open>real_one_collect\<close>, (*"m is_num ==> n + m * n = (1 + m) * n"*)
walther@60385
   272
       \<^rule_thm>\<open>real_one_collect_assoc\<close>, (*"m is_num ==> k + (n + m * n) = k + (1 + m) * n"*)
walther@60358
   273
  
Walther@60515
   274
       \<^rule_eval>\<open>plus\<close> (**)(Calc_Binop.numeric "#add_"), 
Walther@60515
   275
       \<^rule_eval>\<open>minus\<close> (**)(Calc_Binop.numeric "#sub_"), 
Walther@60515
   276
       \<^rule_eval>\<open>times\<close> (**)(Calc_Binop.numeric "#mult_"),
walther@60358
   277
       \<^rule_eval>\<open>divide\<close> (Prog_Expr.eval_cancel "#divide_e"),
walther@60358
   278
       \<^rule_eval>\<open>sqrt\<close> (eval_sqrt "#sqrt_"),
Walther@60515
   279
       \<^rule_eval>\<open>realpow\<close> (**)(Calc_Binop.numeric "#power_")],
Walther@60586
   280
    program = Rule.Empty_Prog});      
wneuper@59472
   281
\<close>
wenzelm@60289
   282
rule_set_knowledge expand_rootbinoms = expand_rootbinoms
walther@60278
   283
ML \<open>
walther@60278
   284
\<close> ML \<open>
walther@60278
   285
\<close> ML \<open>
walther@60278
   286
\<close>
neuper@37906
   287
end