src/Tools/isac/Knowledge/Rational.thy
author wneuper <walther.neuper@jku.at>
Mon, 12 Jul 2021 17:21:37 +0200
changeset 60319 2edbed71fde6
parent 60318 e6e7a9b9ced7
child 60320 02102eaa2021
permissions -rw-r--r--
Test_Some.thy + rewrite.sml + poly.sml ok: repair poly_of_term
neuper@52105
     1
(* rationals, fractions of multivariate polynomials over the real field
neuper@37906
     2
   author: isac team
neuper@52105
     3
   Copyright (c) isac team 2002, 2013
neuper@37906
     4
   Use is subject to license terms.
neuper@37906
     5
neuper@37906
     6
   depends on Poly (and not on Atools), because 
wneuper@59370
     7
   fractions with _normalised_ polynomials are canceled, added, etc.
neuper@37906
     8
*)
neuper@37906
     9
neuper@48884
    10
theory Rational 
wneuper@59553
    11
imports Poly GCD_Poly_ML
neuper@48884
    12
begin
neuper@37906
    13
walther@59878
    14
section \<open>Constants for evaluation by "Rule.Eval"\<close>
neuper@37906
    15
consts
neuper@37906
    16
walther@60278
    17
  is_expanded    :: "real => bool" ("_ is'_expanded")     (*RL->Poly.thy*)
walther@60278
    18
  is_ratpolyexp  :: "real => bool" ("_ is'_ratpolyexp") 
jan@42300
    19
  get_denominator :: "real => real"
walther@60017
    20
  get_numerator   :: "real => real"           
neuper@37906
    21
wneuper@59472
    22
ML \<open>
neuper@52105
    23
(*.the expression contains + - * ^ / only ?.*)
neuper@52105
    24
fun is_ratpolyexp (Free _) = true
neuper@52105
    25
  | is_ratpolyexp (Const ("Groups.plus_class.plus",_) $ Free _ $ Free _) = true
neuper@52105
    26
  | is_ratpolyexp (Const ("Groups.minus_class.minus",_) $ Free _ $ Free _) = true
neuper@52105
    27
  | is_ratpolyexp (Const ("Groups.times_class.times",_) $ Free _ $ Free _) = true
walther@60275
    28
  | is_ratpolyexp (Const ("Transcendental.powr",_) $ Free _ $ Free _) = true
wneuper@59360
    29
  | is_ratpolyexp (Const ("Rings.divide_class.divide",_) $ Free _ $ Free _) = true
neuper@52105
    30
  | is_ratpolyexp (Const ("Groups.plus_class.plus",_) $ t1 $ t2) = 
neuper@52105
    31
               ((is_ratpolyexp t1) andalso (is_ratpolyexp t2))
neuper@52105
    32
  | is_ratpolyexp (Const ("Groups.minus_class.minus",_) $ t1 $ t2) = 
neuper@52105
    33
               ((is_ratpolyexp t1) andalso (is_ratpolyexp t2))
neuper@52105
    34
  | is_ratpolyexp (Const ("Groups.times_class.times",_) $ t1 $ t2) = 
neuper@52105
    35
               ((is_ratpolyexp t1) andalso (is_ratpolyexp t2))
walther@60275
    36
  | is_ratpolyexp (Const ("Transcendental.powr",_) $ t1 $ t2) = 
neuper@52105
    37
               ((is_ratpolyexp t1) andalso (is_ratpolyexp t2))
wneuper@59360
    38
  | is_ratpolyexp (Const ("Rings.divide_class.divide",_) $ t1 $ t2) = 
neuper@52105
    39
               ((is_ratpolyexp t1) andalso (is_ratpolyexp t2))
neuper@52105
    40
  | is_ratpolyexp _ = false;
neuper@37950
    41
walther@60278
    42
(*("is_ratpolyexp", ("Rational.is_ratpolyexp", eval_is_ratpolyexp ""))*)
neuper@52105
    43
fun eval_is_ratpolyexp (thmid:string) _ 
walther@60278
    44
		       (t as (Const("Rational.is_ratpolyexp", _) $ arg)) thy =
neuper@52105
    45
    if is_ratpolyexp arg
walther@59870
    46
    then SOME (TermC.mk_thmid thmid (UnparseC.term_in_thy thy arg) "", 
wneuper@59390
    47
	         HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True})))
walther@59870
    48
    else SOME (TermC.mk_thmid thmid (UnparseC.term_in_thy thy arg) "", 
wneuper@59390
    49
	         HOLogic.Trueprop $ (TermC.mk_equality (t, @{term False})))
neuper@52105
    50
  | eval_is_ratpolyexp _ _ _ _ = NONE; 
neuper@52105
    51
neuper@52105
    52
(*("get_denominator", ("Rational.get_denominator", eval_get_denominator ""))*)
neuper@52105
    53
fun eval_get_denominator (thmid:string) _ 
neuper@52105
    54
		      (t as Const ("Rational.get_denominator", _) $
walther@59603
    55
              (Const ("Rings.divide_class.divide", _) $ _(*num*) $
neuper@52105
    56
                denom)) thy = 
walther@59870
    57
      SOME (TermC.mk_thmid thmid (UnparseC.term_in_thy thy denom) "", 
wneuper@59390
    58
	            HOLogic.Trueprop $ (TermC.mk_equality (t, denom)))
neuper@52105
    59
  | eval_get_denominator _ _ _ _ = NONE; 
neuper@52105
    60
neuper@52105
    61
(*("get_numerator", ("Rational.get_numerator", eval_get_numerator ""))*)
neuper@52105
    62
fun eval_get_numerator (thmid:string) _ 
neuper@52105
    63
      (t as Const ("Rational.get_numerator", _) $
wneuper@59360
    64
          (Const ("Rings.divide_class.divide", _) $num
neuper@52105
    65
            $denom )) thy = 
walther@59870
    66
    SOME (TermC.mk_thmid thmid (UnparseC.term_in_thy thy num) "", 
wneuper@59390
    67
	    HOLogic.Trueprop $ (TermC.mk_equality (t, num)))
neuper@52105
    68
  | eval_get_numerator _ _ _ _ = NONE; 
wneuper@59472
    69
\<close>
neuper@52105
    70
wneuper@59472
    71
section \<open>Theorems for rewriting\<close>
neuper@52105
    72
neuper@52105
    73
axiomatization (* naming due to Isabelle2002, but not contained in Isabelle2002; 
neuper@52105
    74
                  many thms are due to RL and can be removed with updating the equation solver;
neuper@52105
    75
                  TODO: replace by equivalent thms in recent Isabelle201x *) 
neuper@52105
    76
where
neuper@52105
    77
  mult_cross:      "[| b ~= 0; d ~= 0 |] ==> (a / b = c / d) = (a * d = b * c)" and
neuper@52105
    78
  mult_cross1:     "   b ~= 0            ==> (a / b = c    ) = (a     = b * c)" and
neuper@52105
    79
  mult_cross2:     "           d ~= 0    ==> (a     = c / d) = (a * d =     c)" and
neuper@37950
    80
                  
neuper@52105
    81
  add_minus:       "a + b - b = a"(*RL->Poly.thy*) and
neuper@52105
    82
  add_minus1:      "a - b + b = a"(*RL->Poly.thy*) and
neuper@37950
    83
                  
neuper@52105
    84
  rat_mult:        "a / b * (c / d) = a * c / (b * d)"(*?Isa02*)  and
neuper@52105
    85
  rat_mult2:       "a / b *  c      = a * c /  b     "(*?Isa02*) and
t@42211
    86
neuper@52105
    87
  rat_mult_poly_l: "c is_polyexp ==> c * (a / b) = c * a /  b" and
neuper@52105
    88
  rat_mult_poly_r: "c is_polyexp ==> (a / b) * c = a * c /  b" and
neuper@37906
    89
neuper@37906
    90
(*real_times_divide1_eq .. Isa02*) 
neuper@52105
    91
  real_times_divide_1_eq:  "-1 * (c / d) = -1 * c / d " and
neuper@52105
    92
  real_times_divide_num:   "a is_const ==> a * (c / d) = a * c / d " and
t@42211
    93
neuper@52105
    94
  real_mult_div_cancel2:   "k ~= 0 ==> m * k / (n * k) = m / n" and
neuper@37978
    95
(*real_mult_div_cancel1:   "k ~= 0 ==> k * m / (k * n) = m / n"..Isa02*)
neuper@37906
    96
			  
neuper@52105
    97
  real_divide_divide1:     "y ~= 0 ==> (u / v) / (y / z) = (u / v) * (z / y)" and
neuper@52105
    98
  real_divide_divide1_mg:  "y ~= 0 ==> (u / v) / (y / z) = (u * z) / (y * v)" and
neuper@37978
    99
(*real_divide_divide2_eq:  "x / y / z = x / (y * z)"..Isa02*)
neuper@37906
   100
			  
walther@60242
   101
  rat_power:               "(a / b) \<up> n = (a \<up> n) / (b \<up> n)" and
neuper@37978
   102
neuper@37978
   103
  rat_add:         "[| a is_const; b is_const; c is_const; d is_const |] ==> 
neuper@52105
   104
	           a / c + b / d = (a * d + b * c) / (c * d)" and
neuper@37978
   105
  rat_add_assoc:   "[| a is_const; b is_const; c is_const; d is_const |] ==> 
neuper@52105
   106
	           a / c +(b / d + e) = (a * d + b * c)/(d * c) + e" and
neuper@37978
   107
  rat_add1:        "[| a is_const; b is_const; c is_const |] ==> 
neuper@52105
   108
	           a / c + b / c = (a + b) / c" and
neuper@37978
   109
  rat_add1_assoc:   "[| a is_const; b is_const; c is_const |] ==> 
neuper@52105
   110
	           a / c + (b / c + e) = (a + b) / c + e" and
neuper@37978
   111
  rat_add2:        "[| a is_const; b is_const; c is_const |] ==> 
neuper@52105
   112
	           a / c + b = (a + b * c) / c" and
neuper@37978
   113
  rat_add2_assoc:  "[| a is_const; b is_const; c is_const |] ==> 
neuper@52105
   114
	           a / c + (b + e) = (a + b * c) / c + e" and
neuper@37978
   115
  rat_add3:        "[| a is_const; b is_const; c is_const |] ==> 
neuper@52105
   116
	           a + b / c = (a * c + b) / c" and
neuper@37978
   117
  rat_add3_assoc:   "[| a is_const; b is_const; c is_const |] ==> 
neuper@37950
   118
	           a + (b / c + e) = (a * c + b) / c + e"
neuper@37950
   119
wneuper@59472
   120
section \<open>Cancellation and addition of fractions\<close>
wneuper@59472
   121
subsection \<open>Conversion term <--> poly\<close>
wneuper@59472
   122
subsubsection \<open>Convert a term to the internal representation of a multivariate polynomial\<close>
wneuper@59472
   123
ML \<open>
wneuper@59532
   124
fun monom_of_term vs (c, es) (t as Const _) =
wneuper@59532
   125
    (c, list_update es (find_index (curry op = t) vs) 1)
walther@60319
   126
  | monom_of_term _ (c, es) (t as (Const ("Num.numeral_class.numeral", _) $ _)) =
walther@60319
   127
    (t |> HOLogic.dest_number |> snd |> curry op * c, es) (*several numerals in one monom*)
walther@60319
   128
  | monom_of_term _ (c, es) (t as (Const ("Groups.uminus_class.uminus", _) $ _)) =
walther@60319
   129
    (t |> HOLogic.dest_number |> snd |> curry op * c, es) (*several numerals in one monom*)
walther@60318
   130
  | monom_of_term  vs (c, es) (t as Free _) =
walther@60318
   131
    (c, list_update es (find_index (curry op = t) vs) 1)
walther@60319
   132
  | monom_of_term  vs (c, es) (Const ("Transcendental.powr", _) $ (b as Free _) $
walther@60319
   133
      (e as Const ("Num.numeral_class.numeral", _) $ _)) =
walther@60319
   134
    (c, list_update es (find_index (curry op = b) vs) (e |> HOLogic.dest_number |> snd))
walther@60319
   135
  | monom_of_term  vs (c, es) (Const ("Transcendental.powr", _) $ (b as Free _) $
walther@60319
   136
      (e as Const ("Groups.uminus_class.uminus", _) $ _)) =
walther@60319
   137
    (c, list_update es (find_index (curry op = b) vs) (e |> HOLogic.dest_number |> snd))
neuper@52087
   138
  | monom_of_term vs (c, es) (Const ("Groups.times_class.times", _) $ m1 $ m2) =
neuper@52087
   139
    let val (c', es') = monom_of_term vs (c, es) m1
neuper@52087
   140
    in monom_of_term vs (c', es') m2 end
walther@59868
   141
  | monom_of_term _ _ t = raise ERROR ("poly malformed 1 with " ^ UnparseC.term t)
neuper@52087
   142
walther@60318
   143
(*-------v------*)
wneuper@59531
   144
fun monoms_of_term vs (t as Const _) =
wneuper@59531
   145
    [monom_of_term  vs (1, replicate (length vs) 0) t]
walther@60319
   146
  | monoms_of_term vs (t as Const ("Num.numeral_class.numeral", _) $ _) =
walther@60319
   147
    [monom_of_term  vs (1, replicate (length vs) 0) t]
walther@60319
   148
  | monoms_of_term vs (t as Const ("Groups.uminus_class.uminus", _) $ _) =
walther@60319
   149
    [monom_of_term  vs (1, replicate (length vs) 0) t]
wneuper@59531
   150
  | monoms_of_term vs (t as Free _) =
neuper@52087
   151
    [monom_of_term  vs (1, replicate (length vs) 0) t]
walther@60275
   152
  | monoms_of_term vs (t as Const ("Transcendental.powr", _) $ _ $  _) =
neuper@52087
   153
    [monom_of_term  vs (1, replicate (length vs) 0) t]
neuper@52087
   154
  | monoms_of_term vs (t as Const ("Groups.times_class.times", _) $ _ $  _) =
neuper@52087
   155
    [monom_of_term  vs (1, replicate (length vs) 0) t]
neuper@52087
   156
  | monoms_of_term vs (Const ("Groups.plus_class.plus", _) $ ms1 $ ms2) =
neuper@52087
   157
    (monoms_of_term vs ms1) @ (monoms_of_term vs ms2)
walther@59868
   158
  | monoms_of_term _ t = raise ERROR ("poly malformed 2 with " ^ UnparseC.term t)
neuper@52087
   159
neuper@52087
   160
(* convert a term to the internal representation of a multivariate polynomial;
neuper@52087
   161
  the conversion is quite liberal, see test --- fun poly_of_term ---:
neuper@52087
   162
* the order of variables and the parentheses within a monomial are arbitrary
neuper@52087
   163
* the coefficient may be somewhere
neuper@52087
   164
* he order and the parentheses within monomials are arbitrary
neuper@52087
   165
But the term must be completely expand + over * (laws of distributivity are not applicable).
neuper@52087
   166
neuper@52087
   167
The function requires the free variables as strings already given, 
neuper@52087
   168
because the gcd involves 2 polynomials (with the same length for their list of exponents).
neuper@52087
   169
*)
neuper@52087
   170
fun poly_of_term vs (t as Const ("Groups.plus_class.plus", _) $ _ $ _) =
neuper@52101
   171
    (SOME (t |> monoms_of_term vs |> order)
neuper@52087
   172
      handle ERROR _ => NONE)
neuper@52087
   173
  | poly_of_term vs t =
neuper@52101
   174
    (SOME [monom_of_term vs (1, replicate (length vs) 0) t]
neuper@52087
   175
      handle ERROR _ => NONE)
neuper@52087
   176
neuper@52087
   177
fun is_poly t =
wneuper@59532
   178
  let
wneuper@59532
   179
    val vs = TermC.vars_of t
neuper@52087
   180
  in 
neuper@52087
   181
    case poly_of_term vs t of SOME _ => true | NONE => false
neuper@52087
   182
  end
neuper@52105
   183
val is_expanded = is_poly   (* TODO: check names *)
neuper@52105
   184
val is_polynomial = is_poly (* TODO: check names *)
wneuper@59472
   185
\<close>
neuper@52087
   186
wneuper@59472
   187
subsubsection \<open>Convert internal representation of a multivariate polynomial to a term\<close>
wneuper@59472
   188
ML \<open>
neuper@52087
   189
fun term_of_es _ _ _ [] = [] (*assumes same length for vs and es*)
wneuper@59532
   190
  | term_of_es baseT expT (_ :: vs) (0 :: es) = [] @ term_of_es baseT expT vs es
wneuper@59532
   191
  | term_of_es baseT expT (v :: vs) (1 :: es) = v :: term_of_es baseT expT vs es
neuper@52087
   192
  | term_of_es baseT expT (v :: vs) (e :: es) =
walther@60317
   193
    Const ("Transcendental.powr", [baseT, expT] ---> baseT) $ v $ (HOLogic.mk_number expT e)
wneuper@59532
   194
    :: term_of_es baseT expT vs es
wneuper@59532
   195
  | term_of_es _ _ _ _ = raise ERROR "term_of_es: length vs <> length es"
neuper@52087
   196
neuper@52087
   197
fun term_of_monom baseT expT vs ((c, es): monom) =
neuper@52087
   198
    let val es' = term_of_es baseT expT vs es
neuper@52087
   199
    in 
neuper@52087
   200
      if c = 1 
neuper@52087
   201
      then 
neuper@52087
   202
        if es' = [] (*if es = [0,0,0,...]*)
walther@60317
   203
        then HOLogic.mk_number baseT c
neuper@52087
   204
        else foldl (HOLogic.mk_binop "Groups.times_class.times") (hd es', tl es')
walther@60317
   205
      else foldl (HOLogic.mk_binop "Groups.times_class.times")
walther@60317
   206
        (HOLogic.mk_number baseT c, es') 
neuper@52087
   207
    end
neuper@52087
   208
neuper@52087
   209
fun term_of_poly baseT expT vs p =
wneuper@59190
   210
  let val monos = map (term_of_monom baseT expT vs) p
neuper@52087
   211
  in foldl (HOLogic.mk_binop "Groups.plus_class.plus") (hd monos, tl monos) end
wneuper@59472
   212
\<close>
neuper@52087
   213
wneuper@59472
   214
subsection \<open>Apply gcd_poly for cancelling and adding fractions as terms\<close>
wneuper@59472
   215
ML \<open>
neuper@52091
   216
fun mk_noteq_0 baseT t = 
neuper@52091
   217
  Const ("HOL.Not", HOLogic.boolT --> HOLogic.boolT) $ 
walther@60317
   218
    (Const ("HOL.eq", [baseT, baseT] ---> HOLogic.boolT) $ t $ HOLogic.mk_number HOLogic.realT 0)
neuper@52091
   219
neuper@52094
   220
fun mk_asms baseT ts =
wneuper@59389
   221
  let val as' = filter_out TermC.is_num ts (* asm like "2 ~= 0" is needless *)
neuper@52094
   222
  in map (mk_noteq_0 baseT) as' end
wneuper@59472
   223
\<close>
neuper@52094
   224
wneuper@59472
   225
subsubsection \<open>Factor out gcd for cancellation\<close>
wneuper@59472
   226
ML \<open>
neuper@52091
   227
fun check_fraction t =
walther@60269
   228
  case t of
walther@60269
   229
    Const ("Rings.divide_class.divide", _) $ numerator $ denominator
walther@60269
   230
      => SOME (numerator, denominator)
walther@60269
   231
  | _ => NONE
neuper@52091
   232
neuper@52091
   233
(* prepare a term for cancellation by factoring out the gcd
neuper@52091
   234
  assumes: is a fraction with outmost "/"*)
neuper@52091
   235
fun factout_p_ (thy: theory) t =
neuper@52091
   236
  let val opt = check_fraction t
neuper@52091
   237
  in
neuper@52091
   238
    case opt of 
neuper@52091
   239
      NONE => NONE
neuper@52091
   240
    | SOME (numerator, denominator) =>
wneuper@59532
   241
      let
wneuper@59532
   242
        val vs = TermC.vars_of t
neuper@52091
   243
        val baseT = type_of numerator
neuper@52091
   244
        val expT = HOLogic.realT
neuper@52091
   245
      in
neuper@52091
   246
        case (poly_of_term vs numerator, poly_of_term vs denominator) of
neuper@52091
   247
          (SOME a, SOME b) =>
neuper@52091
   248
            let
neuper@52091
   249
              val ((a', b'), c) = gcd_poly a b
neuper@52096
   250
              val es = replicate (length vs) 0 
neuper@52096
   251
            in
neuper@52096
   252
              if c = [(1, es)] orelse c = [(~1, es)]
neuper@52096
   253
              then NONE
neuper@52096
   254
              else 
neuper@52096
   255
                let
neuper@52096
   256
                  val b't = term_of_poly baseT expT vs b'
neuper@52096
   257
                  val ct = term_of_poly baseT expT vs c
neuper@52096
   258
                  val t' = 
wneuper@59360
   259
                    HOLogic.mk_binop "Rings.divide_class.divide" 
neuper@52096
   260
                      (HOLogic.mk_binop "Groups.times_class.times"
wneuper@59190
   261
                        (term_of_poly baseT expT vs a', ct),
neuper@52096
   262
                       HOLogic.mk_binop "Groups.times_class.times" (b't, ct))
neuper@52105
   263
                in SOME (t', mk_asms baseT [b't, ct]) end
neuper@52096
   264
            end
neuper@52091
   265
        | _ => NONE : (term * term list) option
neuper@52091
   266
      end
neuper@52091
   267
  end
wneuper@59472
   268
\<close>
neuper@52091
   269
wneuper@59472
   270
subsubsection \<open>Cancel a fraction\<close>
wneuper@59472
   271
ML \<open>
neuper@52096
   272
(* cancel a term by the gcd ("" denote terms with internal algebraic structure)
neuper@52096
   273
  cancel_p_ :: theory \<Rightarrow> term  \<Rightarrow> (term \<times> term list) option
neuper@52096
   274
  cancel_p_ thy "a / b" = SOME ("a' / b'", ["b' \<noteq> 0"])
neuper@52096
   275
  assumes: a is_polynomial  \<and>  b is_polynomial  \<and>  b \<noteq> 0
neuper@52096
   276
  yields
neuper@52096
   277
    SOME ("a' / b'", ["b' \<noteq> 0"]). gcd_poly a b \<noteq> 1  \<and>  gcd_poly a b \<noteq> -1  \<and>  
neuper@52096
   278
      a' * gcd_poly a b = a  \<and>  b' * gcd_poly a b = b
neuper@52096
   279
    \<or> NONE *)
neuper@52101
   280
fun cancel_p_ (_: theory) t =
neuper@52091
   281
  let val opt = check_fraction t
neuper@52091
   282
  in
neuper@52091
   283
    case opt of 
neuper@52091
   284
      NONE => NONE
neuper@52091
   285
    | SOME (numerator, denominator) =>
wneuper@59532
   286
      let
wneuper@59532
   287
        val vs = TermC.vars_of t
neuper@52091
   288
        val baseT = type_of numerator
neuper@52091
   289
        val expT = HOLogic.realT
neuper@52091
   290
      in
neuper@52091
   291
        case (poly_of_term vs numerator, poly_of_term vs denominator) of
neuper@52091
   292
          (SOME a, SOME b) =>
neuper@52091
   293
            let
neuper@52091
   294
              val ((a', b'), c) = gcd_poly a b
neuper@52096
   295
              val es = replicate (length vs) 0 
neuper@52096
   296
            in
neuper@52096
   297
              if c = [(1, es)] orelse c = [(~1, es)]
neuper@52096
   298
              then NONE
neuper@52096
   299
              else 
neuper@52096
   300
                let
neuper@52096
   301
                  val bt' = term_of_poly baseT expT vs b'
neuper@52096
   302
                  val ct = term_of_poly baseT expT vs c
neuper@52096
   303
                  val t' = 
wneuper@59360
   304
                    HOLogic.mk_binop "Rings.divide_class.divide" 
wneuper@59190
   305
                      (term_of_poly baseT expT vs a', bt')
neuper@52096
   306
                  val asm = mk_asms baseT [bt']
neuper@52096
   307
                in SOME (t', asm) end
neuper@52096
   308
            end
neuper@52091
   309
        | _ => NONE : (term * term list) option
neuper@52091
   310
      end
neuper@52091
   311
  end
wneuper@59472
   312
\<close>
neuper@52091
   313
wneuper@59472
   314
subsubsection \<open>Factor out to a common denominator for addition\<close>
wneuper@59472
   315
ML \<open>
neuper@52101
   316
(* addition of fractions allows (at most) one non-fraction (a monomial) *)
neuper@52101
   317
fun check_frac_sum 
neuper@52091
   318
    (Const ("Groups.plus_class.plus", _) $ 
wneuper@59360
   319
      (Const ("Rings.divide_class.divide", _) $ n1 $ d1) $
wneuper@59360
   320
      (Const ("Rings.divide_class.divide", _) $ n2 $ d2))
neuper@52091
   321
    = SOME ((n1, d1), (n2, d2))
neuper@52101
   322
  | check_frac_sum 
neuper@52091
   323
    (Const ("Groups.plus_class.plus", _) $ 
neuper@52091
   324
      nofrac $ 
wneuper@59360
   325
      (Const ("Rings.divide_class.divide", _) $ n2 $ d2))
walther@60317
   326
    = SOME ((nofrac, HOLogic.mk_number HOLogic.realT 1), (n2, d2))
neuper@52101
   327
  | check_frac_sum 
neuper@52091
   328
    (Const ("Groups.plus_class.plus", _) $ 
wneuper@59360
   329
      (Const ("Rings.divide_class.divide", _) $ n1 $ d1) $ 
neuper@52091
   330
      nofrac)
walther@60317
   331
    = SOME ((n1, d1), (nofrac, HOLogic.mk_number HOLogic.realT 1))
neuper@52101
   332
  | check_frac_sum _ = NONE  
neuper@52091
   333
neuper@52091
   334
(* prepare a term for addition by providing the least common denominator as a product
neuper@52091
   335
  assumes: is a term with outmost "+" and at least one outmost "/" in respective summands*)
neuper@52101
   336
fun common_nominator_p_ (_: theory) t =
neuper@52101
   337
  let val opt = check_frac_sum t
neuper@52091
   338
  in
neuper@52091
   339
    case opt of 
neuper@52091
   340
      NONE => NONE
neuper@52091
   341
    | SOME ((n1, d1), (n2, d2)) =>
wneuper@59532
   342
      let
wneuper@59532
   343
        val vs = TermC.vars_of t
neuper@52091
   344
      in
neuper@52091
   345
        case (poly_of_term vs d1, poly_of_term vs d2) of
neuper@52091
   346
          (SOME a, SOME b) =>
neuper@52091
   347
            let
neuper@52091
   348
              val ((a', b'), c) = gcd_poly a b
neuper@52101
   349
              val (baseT, expT) = (type_of n1, HOLogic.realT)
wneuper@59190
   350
              val [d1', d2', c'] = map (term_of_poly baseT expT vs) [a', b', c]
neuper@52091
   351
              (*----- minimum of parentheses & nice result, but breaks tests: -------------
neuper@52091
   352
              val denom = HOLogic.mk_binop "Groups.times_class.times" 
neuper@52101
   353
                (HOLogic.mk_binop "Groups.times_class.times" (d1', d2'), c') -------------*)
neuper@52101
   354
              val denom =
neuper@52101
   355
                if c = [(1, replicate (length vs) 0)]
neuper@52101
   356
                then HOLogic.mk_binop "Groups.times_class.times" (d1', d2')
neuper@52101
   357
                else
neuper@52101
   358
                  HOLogic.mk_binop "Groups.times_class.times" (c',
neuper@52101
   359
                  HOLogic.mk_binop "Groups.times_class.times" (d1', d2')) (*--------------*)
neuper@52091
   360
              val t' =
neuper@52091
   361
                HOLogic.mk_binop "Groups.plus_class.plus"
wneuper@59360
   362
                  (HOLogic.mk_binop "Rings.divide_class.divide"
neuper@52091
   363
                    (HOLogic.mk_binop "Groups.times_class.times" (n1, d2'), denom),
wneuper@59360
   364
                  HOLogic.mk_binop "Rings.divide_class.divide" 
neuper@52091
   365
                    (HOLogic.mk_binop "Groups.times_class.times" (n2, d1'), denom))
neuper@52094
   366
              val asm = mk_asms baseT [d1', d2', c']
neuper@52091
   367
            in SOME (t', asm) end
neuper@52091
   368
        | _ => NONE : (term * term list) option
neuper@52091
   369
      end
neuper@52091
   370
  end
wneuper@59472
   371
\<close>
neuper@52105
   372
wneuper@59472
   373
subsubsection \<open>Addition of at least one fraction within a sum\<close>
wneuper@59472
   374
ML \<open>
neuper@52091
   375
(* add fractions
neuper@52100
   376
  assumes: is a term with outmost "+" and at least one outmost "/" in respective summands
neuper@52100
   377
  NOTE: the case "(_ + _) + _" need not be considered due to iterated addition.*)
neuper@52105
   378
fun add_fraction_p_ (_: theory) t =
neuper@52101
   379
  case check_frac_sum t of 
neuper@52101
   380
    NONE => NONE
neuper@52101
   381
  | SOME ((n1, d1), (n2, d2)) =>
wneuper@59532
   382
    let
wneuper@59532
   383
      val vs = TermC.vars_of t
neuper@52101
   384
    in
neuper@52101
   385
      case (poly_of_term vs n1, poly_of_term vs d1, poly_of_term vs n2, poly_of_term vs d2) of
neuper@52101
   386
        (SOME _, SOME a, SOME _, SOME b) =>
neuper@52101
   387
          let
neuper@52101
   388
            val ((a', b'), c) = gcd_poly a b
neuper@52101
   389
            val (baseT, expT) = (type_of n1, HOLogic.realT)
neuper@52101
   390
            val nomin = term_of_poly baseT expT vs 
neuper@52101
   391
              (((the (poly_of_term vs n1)) %%*%% b') %%+%% ((the (poly_of_term vs n2)) %%*%% a')) 
neuper@52101
   392
            val denom = term_of_poly baseT expT vs ((c %%*%% a') %%*%% b')
wneuper@59360
   393
            val t' = HOLogic.mk_binop "Rings.divide_class.divide" (nomin, denom)
neuper@52101
   394
          in SOME (t', mk_asms baseT [denom]) end
neuper@52101
   395
      | _ => NONE : (term * term list) option
neuper@52101
   396
    end
wneuper@59472
   397
\<close>
neuper@52091
   398
wneuper@59472
   399
section \<open>Embed cancellation and addition into rewriting\<close>
wneuper@59472
   400
ML \<open>val thy = @{theory}\<close>
wneuper@59472
   401
subsection \<open>Rulesets and predicate for embedding\<close>
wneuper@59472
   402
ML \<open>
neuper@52105
   403
(* evaluates conditions in calculate_Rational *)
neuper@52105
   404
val calc_rat_erls =
s1210629013@55444
   405
  prep_rls'
walther@59857
   406
    (Rule_Def.Repeat {id = "calc_rat_erls", preconds = [], rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord), 
walther@59852
   407
      erls = Rule_Set.empty, srls = Rule_Set.Empty, calc = [], errpatts = [],
neuper@52105
   408
      rules = 
walther@59878
   409
        [Rule.Eval ("HOL.eq", Prog_Expr.eval_equal "#equal_"),
walther@60278
   410
        Rule.Eval ("Prog_Expr.is_const", Prog_Expr.eval_const "#is_const_"),
walther@59871
   411
        Rule.Thm ("not_true", ThmC.numerals_to_Free @{thm not_true}),
walther@59871
   412
        Rule.Thm ("not_false", ThmC.numerals_to_Free @{thm not_false})], 
walther@59878
   413
      scr = Rule.Empty_Prog});
neuper@37950
   414
neuper@52105
   415
(* simplifies expressions with numerals;
neuper@52105
   416
   does NOT rearrange the term by AC-rewriting; thus terms with variables 
neuper@52105
   417
   need to have constants to be commuted together respectively           *)
neuper@52105
   418
val calculate_Rational =
walther@59852
   419
  prep_rls' (Rule_Set.merge "calculate_Rational"
walther@59857
   420
    (Rule_Def.Repeat {id = "divide", preconds = [], rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord), 
walther@59851
   421
      erls = calc_rat_erls, srls = Rule_Set.Empty,
neuper@52105
   422
      calc = [], errpatts = [],
neuper@52105
   423
      rules = 
walther@59878
   424
        [Rule.Eval ("Rings.divide_class.divide", Prog_Expr.eval_cancel "#divide_e"),
neuper@37950
   425
walther@59871
   426
        Rule.Thm ("minus_divide_left", ThmC.numerals_to_Free (@{thm minus_divide_left} RS @{thm sym})),
neuper@52105
   427
          (*SYM - ?x / ?y = - (?x / ?y)  may come from subst*)
walther@59871
   428
        Rule.Thm ("rat_add", ThmC.numerals_to_Free @{thm rat_add}),
neuper@52105
   429
          (*"[| a is_const; b is_const; c is_const; d is_const |] ==> \
neuper@52105
   430
          \a / c + b / d = (a * d) / (c * d) + (b * c ) / (d * c)"*)
walther@59871
   431
        Rule.Thm ("rat_add1", ThmC.numerals_to_Free @{thm rat_add1}),
neuper@52105
   432
          (*"[| a is_const; b is_const; c is_const |] ==> a / c + b / c = (a + b) / c"*)
walther@59871
   433
        Rule.Thm ("rat_add2", ThmC.numerals_to_Free @{thm rat_add2}),
neuper@52105
   434
          (*"[| ?a is_const; ?b is_const; ?c is_const |] ==> ?a / ?c + ?b = (?a + ?b * ?c) / ?c"*)
walther@59871
   435
        Rule.Thm ("rat_add3", ThmC.numerals_to_Free @{thm rat_add3}),
neuper@52105
   436
          (*"[| a is_const; b is_const; c is_const |] ==> a + b / c = (a * c) / c + b / c"\
neuper@52105
   437
          .... is_const to be omitted here FIXME*)
neuper@52105
   438
        
walther@59871
   439
        Rule.Thm ("rat_mult", ThmC.numerals_to_Free @{thm rat_mult}), 
neuper@52105
   440
          (*a / b * (c / d) = a * c / (b * d)*)
walther@59871
   441
        Rule.Thm ("times_divide_eq_right", ThmC.numerals_to_Free @{thm times_divide_eq_right}),
neuper@52105
   442
          (*?x * (?y / ?z) = ?x * ?y / ?z*)
walther@59871
   443
        Rule.Thm ("times_divide_eq_left", ThmC.numerals_to_Free @{thm times_divide_eq_left}),
neuper@52105
   444
          (*?y / ?z * ?x = ?y * ?x / ?z*)
neuper@52105
   445
        
walther@59871
   446
        Rule.Thm ("real_divide_divide1", ThmC.numerals_to_Free @{thm real_divide_divide1}),
neuper@52105
   447
          (*"?y ~= 0 ==> ?u / ?v / (?y / ?z) = ?u / ?v * (?z / ?y)"*)
walther@59871
   448
        Rule.Thm ("divide_divide_eq_left", ThmC.numerals_to_Free @{thm divide_divide_eq_left}),
neuper@52105
   449
          (*"?x / ?y / ?z = ?x / (?y * ?z)"*)
neuper@52105
   450
        
walther@59871
   451
        Rule.Thm ("rat_power", ThmC.numerals_to_Free @{thm rat_power}),
walther@60260
   452
          (*"(?a / ?b)  \<up>  ?n = ?a  \<up>  ?n / ?b  \<up>  ?n"*)
neuper@52105
   453
        
walther@59871
   454
        Rule.Thm ("mult_cross", ThmC.numerals_to_Free @{thm mult_cross}),
neuper@52105
   455
          (*"[| b ~= 0; d ~= 0 |] ==> (a / b = c / d) = (a * d = b * c)*)
walther@59871
   456
        Rule.Thm ("mult_cross1", ThmC.numerals_to_Free @{thm mult_cross1}),
neuper@52105
   457
          (*"   b ~= 0            ==> (a / b = c    ) = (a     = b * c)*)
walther@59871
   458
        Rule.Thm ("mult_cross2", ThmC.numerals_to_Free @{thm mult_cross2})
neuper@52105
   459
          (*"           d ~= 0    ==> (a     = c / d) = (a * d =     c)*)], 
walther@59878
   460
      scr = Rule.Empty_Prog})
neuper@52105
   461
    calculate_Poly);
neuper@37950
   462
walther@60278
   463
(*("is_expanded", ("Rational.is_expanded", eval_is_expanded ""))*)
neuper@37950
   464
fun eval_is_expanded (thmid:string) _ 
walther@60278
   465
		       (t as (Const("Rational.is_expanded", _) $ arg)) thy = 
neuper@37950
   466
    if is_expanded arg
walther@59870
   467
    then SOME (TermC.mk_thmid thmid (UnparseC.term_in_thy thy arg) "", 
wneuper@59390
   468
	         HOLogic.Trueprop $ (TermC.mk_equality (t, @{term True})))
walther@59870
   469
    else SOME (TermC.mk_thmid thmid (UnparseC.term_in_thy thy arg) "", 
wneuper@59390
   470
	         HOLogic.Trueprop $ (TermC.mk_equality (t, @{term False})))
s1210629013@52159
   471
  | eval_is_expanded _ _ _ _ = NONE;
wneuper@59472
   472
\<close>
wneuper@59472
   473
setup \<open>KEStore_Elems.add_calcs
walther@60278
   474
  [("is_expanded", ("Rational.is_expanded", eval_is_expanded ""))]\<close>
wneuper@59472
   475
ML \<open>
neuper@37950
   476
val rational_erls = 
walther@59852
   477
  Rule_Set.merge "rational_erls" calculate_Rational 
walther@59852
   478
    (Rule_Set.append_rules "is_expanded" Atools_erls 
walther@60278
   479
      [Rule.Eval ("Rational.is_expanded", eval_is_expanded "")]);
wneuper@59472
   480
\<close>
neuper@37950
   481
wneuper@59472
   482
subsection \<open>Embed cancellation into rewriting\<close>
wneuper@59472
   483
ML \<open>
walther@59603
   484
(**)local (* cancel_p *)
neuper@37950
   485
walther@59852
   486
val {rules = rules, rew_ord = (_, ro), ...} = Rule_Set.rep (assoc_rls' @{theory} "rev_rew_p");
neuper@37950
   487
neuper@52105
   488
fun init_state thy eval_rls ro t =
neuper@52105
   489
  let
neuper@52105
   490
    val SOME (t', _) = factout_p_ thy t;
neuper@52105
   491
    val SOME (t'', asm) = cancel_p_ thy t;
walther@59907
   492
    val der = Derive.steps_reverse thy eval_rls rules ro NONE t';
neuper@52105
   493
    val der = der @ 
walther@59871
   494
      [(Rule.Thm ("real_mult_div_cancel2", ThmC.numerals_to_Free @{thm real_mult_div_cancel2}), (t'', asm))]
walther@60017
   495
    val rs = (Rule.distinct' o (map #1)) der
walther@60017
   496
  	val rs = filter_out (ThmC.member'
neuper@52105
   497
  	  ["sym_real_add_zero_left", "sym_real_mult_0", "sym_real_mult_1"]) rs
neuper@52105
   498
  in (t, t'', [rs(*one in order to ease locate_rule*)], der) end;
neuper@37950
   499
neuper@52105
   500
fun locate_rule thy eval_rls ro [rs] t r =
walther@59876
   501
    if member op = ((map (Rule.thm_id)) rs) (Rule.thm_id r)
neuper@52105
   502
    then 
walther@59876
   503
      let val ropt = Rewrite.rewrite_ thy ro eval_rls true (Rule.thm r) t;
neuper@52105
   504
      in
neuper@52105
   505
        case ropt of SOME ta => [(r, ta)]
walther@59733
   506
	      | NONE => ((*tracing 
walther@59876
   507
	          ("### locate_rule:  rewrite " ^ Rule.thm_id r ^ " " ^ UnparseC.term t ^ " = NONE");*) []) 
neuper@52105
   508
			end
walther@59876
   509
    else ((*tracing ("### locate_rule:  " ^ Rule.thm_id r ^ " not mem rrls");*) [])
walther@59962
   510
  | locate_rule _ _ _ _ _ _ = raise ERROR "locate_rule: doesnt match rev-sets in istate";
neuper@37950
   511
neuper@52105
   512
fun next_rule thy eval_rls ro [rs] t =
neuper@52105
   513
    let
walther@59907
   514
      val der = Derive.do_one thy eval_rls rs ro NONE t;
neuper@52105
   515
    in case der of (_, r, _) :: _ => SOME r | _ => NONE end
walther@59962
   516
  | next_rule _ _ _ _ _ = raise ERROR ("next_rule: doesnt match rev-sets in istate");
neuper@37950
   517
wneuper@59416
   518
fun attach_form (_: Rule.rule list list) (_: term) (_: term) = 
wneuper@59416
   519
  [(*TODO*)]: ( Rule.rule * (term * term list)) list;
neuper@37950
   520
walther@59861
   521
(**)in(**)
neuper@37950
   522
neuper@52105
   523
val cancel_p = 
walther@59850
   524
  Rule_Set.Rrls {id = "cancel_p", prepat = [],
neuper@52105
   525
	rew_ord=("ord_make_polynomial", ord_make_polynomial false thy),
neuper@52105
   526
	erls = rational_erls, 
neuper@52105
   527
	calc = 
walther@59603
   528
	  [("PLUS", ("Groups.plus_class.plus", (**)eval_binop "#add_")),
walther@59603
   529
	  ("TIMES" , ("Groups.times_class.times", (**)eval_binop "#mult_")),
walther@59603
   530
	  ("DIVIDE", ("Rings.divide_class.divide", Prog_Expr.eval_cancel "#divide_e")),
walther@60275
   531
	  ("POWER", ("Transcendental.powr", (**)eval_binop "#power_"))],
neuper@52105
   532
    errpatts = [],
neuper@52105
   533
	scr =
wneuper@59416
   534
	  Rule.Rfuns {init_state  = init_state thy Atools_erls ro,
neuper@52105
   535
		normal_form = cancel_p_ thy, 
neuper@52105
   536
		locate_rule = locate_rule thy Atools_erls ro,
neuper@52105
   537
		next_rule   = next_rule thy Atools_erls ro,
neuper@52105
   538
		attach_form = attach_form}}
walther@59861
   539
(**)end(* local cancel_p *)
wneuper@59472
   540
\<close>
neuper@37950
   541
wneuper@59472
   542
subsection \<open>Embed addition into rewriting\<close>
wneuper@59472
   543
ML \<open>
walther@59861
   544
(**)local (* add_fractions_p *)
neuper@37950
   545
walther@59852
   546
(*val {rules = rules, rew_ord = (_, ro), ...} = Rule_Set.rep (assoc_rls "make_polynomial");*)
walther@59852
   547
val {rules, rew_ord=(_,ro),...} = Rule_Set.rep (assoc_rls' @{theory} "rev_rew_p");
neuper@37950
   548
neuper@52105
   549
fun init_state thy eval_rls ro t =
neuper@52105
   550
  let 
neuper@52105
   551
    val SOME (t',_) = common_nominator_p_ thy t;
neuper@52105
   552
    val SOME (t'', asm) = add_fraction_p_ thy t;
walther@59907
   553
    val der = Derive.steps_reverse thy eval_rls rules ro NONE t';
neuper@52105
   554
    val der = der @ 
walther@59871
   555
      [(Rule.Thm ("real_mult_div_cancel2", ThmC.numerals_to_Free @{thm real_mult_div_cancel2}), (t'',asm))]
walther@60017
   556
    val rs = (Rule.distinct' o (map #1)) der;
walther@60017
   557
    val rs = filter_out (ThmC.member'
neuper@52105
   558
      ["sym_real_add_zero_left", "sym_real_mult_0", "sym_real_mult_1"]) rs;
neuper@52105
   559
  in (t, t'', [rs(*here only _ONE_*)], der) end;
neuper@37950
   560
neuper@52105
   561
fun locate_rule thy eval_rls ro [rs] t r =
walther@59876
   562
    if member op = ((map (Rule.thm_id)) rs) (Rule.thm_id r)
neuper@52105
   563
    then 
walther@59876
   564
      let val ropt = Rewrite.rewrite_ thy ro eval_rls true (Rule.thm r) t;
neuper@52105
   565
      in 
neuper@52105
   566
        case ropt of
neuper@52105
   567
          SOME ta => [(r, ta)]
neuper@52105
   568
	      | NONE => 
walther@59876
   569
	        ((*tracing ("### locate_rule:  rewrite " ^ Rule.thm_id r ^ " " ^ UnparseC.term t ^ " = NONE");*)
neuper@52105
   570
	        []) end
walther@59876
   571
    else ((*tracing ("### locate_rule:  " ^ Rule.thm_id r ^ " not mem rrls");*) [])
walther@59962
   572
  | locate_rule _ _ _ _ _ _ = raise ERROR "locate_rule: doesnt match rev-sets in istate";
neuper@37950
   573
neuper@37950
   574
fun next_rule thy eval_rls ro [rs] t =
walther@59907
   575
    let val der = Derive.do_one thy eval_rls rs ro NONE t;
neuper@52105
   576
    in 
neuper@52105
   577
      case der of
neuper@52105
   578
	      (_,r,_)::_ => SOME r
neuper@52105
   579
	    | _ => NONE
neuper@37950
   580
    end
walther@59962
   581
  | next_rule _ _ _ _ _ = raise ERROR ("next_rule: doesnt match rev-sets in istate");
neuper@37950
   582
wneuper@59389
   583
val pat0 = TermC.parse_patt thy "?r/?s+?u/?v :: real";
wneuper@59389
   584
val pat1 = TermC.parse_patt thy "?r/?s+?u    :: real";
wneuper@59389
   585
val pat2 = TermC.parse_patt thy "?r   +?u/?v :: real";
neuper@48760
   586
val prepat = [([@{term True}], pat0),
neuper@48760
   587
	      ([@{term True}], pat1),
neuper@48760
   588
	      ([@{term True}], pat2)];
walther@59861
   589
(**)in(**)
neuper@37950
   590
neuper@52105
   591
val add_fractions_p =
walther@59850
   592
  Rule_Set.Rrls {id = "add_fractions_p", prepat=prepat,
neuper@52105
   593
    rew_ord = ("ord_make_polynomial", ord_make_polynomial false thy),
neuper@52105
   594
    erls = rational_erls,
walther@59603
   595
    calc = [("PLUS", ("Groups.plus_class.plus", (**)eval_binop "#add_")),
walther@59603
   596
      ("TIMES", ("Groups.times_class.times", (**)eval_binop "#mult_")),
walther@59603
   597
      ("DIVIDE", ("Rings.divide_class.divide", Prog_Expr.eval_cancel "#divide_e")),
walther@60275
   598
      ("POWER", ("Transcendental.powr", (**)eval_binop "#power_"))],
neuper@52105
   599
    errpatts = [],
wneuper@59416
   600
    scr = Rule.Rfuns {init_state  = init_state thy Atools_erls ro,
neuper@52105
   601
      normal_form = add_fraction_p_ thy,
neuper@52105
   602
      locate_rule = locate_rule thy Atools_erls ro,
neuper@52105
   603
      next_rule   = next_rule thy Atools_erls ro,
neuper@52105
   604
      attach_form = attach_form}}
walther@59861
   605
(**)end(*local add_fractions_p *)
wneuper@59472
   606
\<close>
neuper@37950
   607
wneuper@59472
   608
subsection \<open>Cancelling and adding all occurrences in a term /////////////////////////////\<close>
wneuper@59472
   609
ML \<open>
neuper@52105
   610
(*copying cancel_p_rls + add her caused error in interface.sml*)
wneuper@59472
   611
\<close>
neuper@42451
   612
wneuper@59472
   613
section \<open>Rulesets for general simplification\<close>
wneuper@59472
   614
ML \<open>
neuper@37950
   615
(*.all powers over + distributed; atoms over * collected, other distributed
neuper@37950
   616
   contains absolute minimum of thms for context in norm_Rational .*)
s1210629013@55444
   617
val powers = prep_rls'(
walther@59857
   618
  Rule_Def.Repeat {id = "powers", preconds = [], rew_ord = ("dummy_ord",Rewrite_Ord.dummy_ord), 
walther@59851
   619
      erls = powers_erls, srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@59871
   620
      rules = [Rule.Thm ("realpow_multI", ThmC.numerals_to_Free @{thm realpow_multI}),
walther@60260
   621
	       (*"(r * s)  \<up>  n = r  \<up>  n * s  \<up>  n"*)
walther@59871
   622
	       Rule.Thm ("realpow_pow",ThmC.numerals_to_Free @{thm realpow_pow}),
walther@60260
   623
	       (*"(a  \<up>  b)  \<up>  c = a  \<up>  (b * c)"*)
walther@59871
   624
	       Rule.Thm ("realpow_oneI",ThmC.numerals_to_Free @{thm realpow_oneI}),
walther@60260
   625
	       (*"r  \<up>  1 = r"*)
walther@59871
   626
	       Rule.Thm ("realpow_minus_even",ThmC.numerals_to_Free @{thm realpow_minus_even}),
walther@60260
   627
	       (*"n is_even ==> (- r)  \<up>  n = r  \<up>  n" ?-->discard_minus?*)
walther@59871
   628
	       Rule.Thm ("realpow_minus_odd",ThmC.numerals_to_Free @{thm realpow_minus_odd}),
walther@60260
   629
	       (*"Not (n is_even) ==> (- r)  \<up>  n = -1 * r  \<up>  n"*)
neuper@37950
   630
	       
neuper@37950
   631
	       (*----- collect atoms over * -----*)
walther@59871
   632
	       Rule.Thm ("realpow_two_atom",ThmC.numerals_to_Free @{thm realpow_two_atom}),	
walther@60260
   633
	       (*"r is_atom ==> r * r = r  \<up>  2"*)
walther@59871
   634
	       Rule.Thm ("realpow_plus_1",ThmC.numerals_to_Free @{thm realpow_plus_1}),		
walther@60260
   635
	       (*"r is_atom ==> r * r  \<up>  n = r  \<up>  (n + 1)"*)
walther@59871
   636
	       Rule.Thm ("realpow_addI_atom",ThmC.numerals_to_Free @{thm realpow_addI_atom}),
walther@60260
   637
	       (*"r is_atom ==> r  \<up>  n * r  \<up>  m = r  \<up>  (n + m)"*)
neuper@37950
   638
neuper@37950
   639
	       (*----- distribute none-atoms -----*)
walther@59871
   640
	       Rule.Thm ("realpow_def_atom",ThmC.numerals_to_Free @{thm realpow_def_atom}),
walther@60276
   641
	       (*"[| 1 < n; ~ (r is_atom) |]==>r  \<up>  n = r * r  \<up>  (n + -1)"*)
walther@59871
   642
	       Rule.Thm ("realpow_eq_oneI",ThmC.numerals_to_Free @{thm realpow_eq_oneI}),
walther@60260
   643
	       (*"1  \<up>  n = 1"*)
walther@59878
   644
	       Rule.Eval ("Groups.plus_class.plus", (**)eval_binop "#add_")
neuper@37950
   645
	       ],
walther@59878
   646
      scr = Rule.Empty_Prog
wneuper@59406
   647
      });
neuper@37950
   648
(*.contains absolute minimum of thms for context in norm_Rational.*)
s1210629013@55444
   649
val rat_mult_divide = prep_rls'(
walther@59851
   650
  Rule_Def.Repeat {id = "rat_mult_divide", preconds = [], 
walther@59857
   651
      rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord), 
walther@59852
   652
      erls = Rule_Set.empty, srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@59871
   653
      rules = [Rule.Thm ("rat_mult",ThmC.numerals_to_Free @{thm rat_mult}),
neuper@37950
   654
	       (*(1)"?a / ?b * (?c / ?d) = ?a * ?c / (?b * ?d)"*)
walther@59871
   655
	       Rule.Thm ("times_divide_eq_right",ThmC.numerals_to_Free @{thm times_divide_eq_right}),
neuper@37950
   656
	       (*(2)"?a * (?c / ?d) = ?a * ?c / ?d" must be [2],
neuper@37950
   657
	       otherwise inv.to a / b / c = ...*)
walther@59871
   658
	       Rule.Thm ("times_divide_eq_left",ThmC.numerals_to_Free @{thm times_divide_eq_left}),
walther@60260
   659
	       (*"?a / ?b * ?c = ?a * ?c / ?b" order weights x \<up> n too much
walther@60260
   660
		     and does not commute a / b * c  \<up>  2 !*)
neuper@37950
   661
	       
wneuper@59416
   662
	       Rule.Thm ("divide_divide_eq_right", 
walther@59871
   663
                     ThmC.numerals_to_Free @{thm divide_divide_eq_right}),
neuper@37950
   664
	       (*"?x / (?y / ?z) = ?x * ?z / ?y"*)
wneuper@59416
   665
	       Rule.Thm ("divide_divide_eq_left",
walther@59871
   666
                     ThmC.numerals_to_Free @{thm divide_divide_eq_left}),
neuper@37950
   667
	       (*"?x / ?y / ?z = ?x / (?y * ?z)"*)
walther@59878
   668
	       Rule.Eval ("Rings.divide_class.divide", Prog_Expr.eval_cancel "#divide_e")
neuper@37950
   669
	       ],
walther@59878
   670
      scr = Rule.Empty_Prog
wneuper@59406
   671
      });
neuper@37979
   672
neuper@37950
   673
(*.contains absolute minimum of thms for context in norm_Rational.*)
s1210629013@55444
   674
val reduce_0_1_2 = prep_rls'(
walther@59857
   675
  Rule_Def.Repeat{id = "reduce_0_1_2", preconds = [], rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
walther@59852
   676
      erls = Rule_Set.empty, srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@59871
   677
      rules = [(*Rule.Thm ("divide_1",ThmC.numerals_to_Free @{thm divide_1}),
neuper@37950
   678
		 "?x / 1 = ?x" unnecess.for normalform*)
walther@59871
   679
	       Rule.Thm ("mult_1_left",ThmC.numerals_to_Free @{thm mult_1_left}),                 
neuper@37950
   680
	       (*"1 * z = z"*)
walther@59871
   681
	       (*Rule.Thm ("real_mult_minus1",ThmC.numerals_to_Free @{thm real_mult_minus1}),
neuper@37950
   682
	       "-1 * z = - z"*)
walther@59871
   683
	       (*Rule.Thm ("real_minus_mult_cancel",ThmC.numerals_to_Free @{thm real_minus_mult_cancel}),
neuper@37950
   684
	       "- ?x * - ?y = ?x * ?y"*)
neuper@37950
   685
walther@59871
   686
	       Rule.Thm ("mult_zero_left",ThmC.numerals_to_Free @{thm mult_zero_left}),        
neuper@37950
   687
	       (*"0 * z = 0"*)
walther@59871
   688
	       Rule.Thm ("add_0_left",ThmC.numerals_to_Free @{thm add_0_left}),
neuper@37950
   689
	       (*"0 + z = z"*)
walther@59871
   690
	       (*Rule.Thm ("right_minus",ThmC.numerals_to_Free @{thm right_minus}),
neuper@37950
   691
	       "?z + - ?z = 0"*)
neuper@37950
   692
wneuper@59416
   693
	       Rule.Thm ("sym_real_mult_2",
walther@59871
   694
                     ThmC.numerals_to_Free (@{thm real_mult_2} RS @{thm sym})),	
neuper@37950
   695
	       (*"z1 + z1 = 2 * z1"*)
walther@59871
   696
	       Rule.Thm ("real_mult_2_assoc",ThmC.numerals_to_Free @{thm real_mult_2_assoc}),
neuper@37950
   697
	       (*"z1 + (z1 + k) = 2 * z1 + k"*)
neuper@37950
   698
walther@59871
   699
	       Rule.Thm ("division_ring_divide_zero",ThmC.numerals_to_Free @{thm division_ring_divide_zero})
neuper@37950
   700
	       (*"0 / ?x = 0"*)
walther@59878
   701
	       ], scr = Rule.Empty_Prog});
neuper@37950
   702
neuper@37950
   703
(*erls for calculate_Rational; 
neuper@37950
   704
  make local with FIXX@ME result:term *term list WN0609???SKMG*)
s1210629013@55444
   705
val norm_rat_erls = prep_rls'(
walther@59857
   706
  Rule_Def.Repeat {id = "norm_rat_erls", preconds = [], rew_ord = ("dummy_ord",Rewrite_Ord.dummy_ord), 
walther@59852
   707
      erls = Rule_Set.empty, srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@60278
   708
      rules = [Rule.Eval ("Prog_Expr.is_const", Prog_Expr.eval_const "#is_const_")
walther@59878
   709
	       ], scr = Rule.Empty_Prog});
neuper@37979
   710
neuper@52105
   711
(* consists of rls containing the absolute minimum of thms *)
neuper@37950
   712
(*040209: this version has been used by RL for his equations,
neuper@52105
   713
which is now replaced by MGs version "norm_Rational" below *)
s1210629013@55444
   714
val norm_Rational_min = prep_rls'(
walther@59857
   715
  Rule_Def.Repeat {id = "norm_Rational_min", preconds = [], rew_ord = ("dummy_ord",Rewrite_Ord.dummy_ord), 
walther@59851
   716
      erls = norm_rat_erls, srls = Rule_Set.Empty, calc = [], errpatts = [],
neuper@37950
   717
      rules = [(*sequence given by operator precedence*)
wneuper@59416
   718
	       Rule.Rls_ discard_minus,
wneuper@59416
   719
	       Rule.Rls_ powers,
wneuper@59416
   720
	       Rule.Rls_ rat_mult_divide,
wneuper@59416
   721
	       Rule.Rls_ expand,
wneuper@59416
   722
	       Rule.Rls_ reduce_0_1_2,
wneuper@59416
   723
	       Rule.Rls_ order_add_mult,
wneuper@59416
   724
	       Rule.Rls_ collect_numerals,
wneuper@59416
   725
	       Rule.Rls_ add_fractions_p,
wneuper@59416
   726
	       Rule.Rls_ cancel_p
neuper@37950
   727
	       ],
walther@59878
   728
      scr = Rule.Empty_Prog});
neuper@37979
   729
s1210629013@55444
   730
val norm_Rational_parenthesized = prep_rls'(
walther@59878
   731
  Rule_Set.Sequence {id = "norm_Rational_parenthesized", preconds = []:term list, 
walther@59857
   732
       rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord),
walther@59851
   733
      erls = Atools_erls, srls = Rule_Set.Empty,
neuper@42451
   734
      calc = [], errpatts = [],
wneuper@59416
   735
      rules = [Rule.Rls_  norm_Rational_min,
wneuper@59416
   736
	       Rule.Rls_ discard_parentheses
neuper@37950
   737
	       ],
walther@59878
   738
      scr = Rule.Empty_Prog});      
neuper@37950
   739
neuper@37950
   740
(*WN030318???SK: simplifies all but cancel and common_nominator*)
neuper@37950
   741
val simplify_rational = 
walther@59852
   742
    Rule_Set.merge "simplify_rational" expand_binoms
walther@59852
   743
    (Rule_Set.append_rules "divide" calculate_Rational
walther@59871
   744
		[Rule.Thm ("div_by_1",ThmC.numerals_to_Free @{thm div_by_1}),
neuper@37950
   745
		 (*"?x / 1 = ?x"*)
walther@59871
   746
		 Rule.Thm ("rat_mult",ThmC.numerals_to_Free @{thm rat_mult}),
neuper@37950
   747
		 (*(1)"?a / ?b * (?c / ?d) = ?a * ?c / (?b * ?d)"*)
walther@59871
   748
		 Rule.Thm ("times_divide_eq_right",ThmC.numerals_to_Free @{thm times_divide_eq_right}),
neuper@37950
   749
		 (*(2)"?a * (?c / ?d) = ?a * ?c / ?d" must be [2],
neuper@37950
   750
		 otherwise inv.to a / b / c = ...*)
walther@59871
   751
		 Rule.Thm ("times_divide_eq_left",ThmC.numerals_to_Free @{thm times_divide_eq_left}),
neuper@37950
   752
		 (*"?a / ?b * ?c = ?a * ?c / ?b"*)
walther@59871
   753
		 Rule.Thm ("add_minus",ThmC.numerals_to_Free @{thm add_minus}),
neuper@37950
   754
		 (*"?a + ?b - ?b = ?a"*)
walther@59871
   755
		 Rule.Thm ("add_minus1",ThmC.numerals_to_Free @{thm add_minus1}),
neuper@37950
   756
		 (*"?a - ?b + ?b = ?a"*)
walther@59871
   757
		 Rule.Thm ("divide_minus1",ThmC.numerals_to_Free @{thm divide_minus1})
neuper@37950
   758
		 (*"?x / -1 = - ?x"*)
neuper@37950
   759
		 ]);
wneuper@59472
   760
\<close>
wneuper@59472
   761
ML \<open>
s1210629013@55444
   762
val add_fractions_p_rls = prep_rls'(
walther@59857
   763
  Rule_Def.Repeat {id = "add_fractions_p_rls", preconds = [], rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord), 
walther@59852
   764
	  erls = Rule_Set.empty, srls = Rule_Set.Empty, calc = [], errpatts = [],
wneuper@59416
   765
	  rules = [Rule.Rls_ add_fractions_p], 
walther@59878
   766
	  scr = Rule.Empty_Prog});
neuper@37950
   767
walther@59851
   768
(* "Rule_Def.Repeat" causes repeated application of cancel_p to one and the same term *)
s1210629013@55444
   769
val cancel_p_rls = prep_rls'(
walther@59851
   770
  Rule_Def.Repeat 
walther@59857
   771
    {id = "cancel_p_rls", preconds = [], rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord), 
walther@59852
   772
    erls = Rule_Set.empty, srls = Rule_Set.Empty, calc = [], errpatts = [],
wneuper@59416
   773
    rules = [Rule.Rls_ cancel_p], 
walther@59878
   774
	  scr = Rule.Empty_Prog});
neuper@52105
   775
neuper@37950
   776
(*. makes 'normal' fractions; 'is_polyexp' inhibits double fractions;
neuper@37950
   777
    used in initial part norm_Rational_mg, see example DA-M02-main.p.60.*)
s1210629013@55444
   778
val rat_mult_poly = prep_rls'(
walther@59857
   779
  Rule_Def.Repeat {id = "rat_mult_poly", preconds = [], rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord), 
walther@60278
   780
	  erls = Rule_Set.append_rules "Rule_Set.empty-is_polyexp" Rule_Set.empty [Rule.Eval ("Poly.is_polyexp", eval_is_polyexp "")], 
walther@59851
   781
	  srls = Rule_Set.Empty, calc = [], errpatts = [],
neuper@52105
   782
	  rules = 
walther@59871
   783
	    [Rule.Thm ("rat_mult_poly_l",ThmC.numerals_to_Free @{thm rat_mult_poly_l}),
neuper@52105
   784
	    (*"?c is_polyexp ==> ?c * (?a / ?b) = ?c * ?a / ?b"*)
walther@59871
   785
	    Rule.Thm ("rat_mult_poly_r",ThmC.numerals_to_Free @{thm rat_mult_poly_r})
neuper@52105
   786
	    (*"?c is_polyexp ==> ?a / ?b * ?c = ?a * ?c / ?b"*) ], 
walther@59878
   787
	  scr = Rule.Empty_Prog});
neuper@37979
   788
neuper@37950
   789
(*. makes 'normal' fractions; 'is_polyexp' inhibits double fractions;
neuper@37950
   790
    used in looping part norm_Rational_rls, see example DA-M02-main.p.60 
walther@59852
   791
    .. WHERE THE LATTER DOES ALWAYS WORK, BECAUSE erls = Rule_Set.empty, 
wneuper@59416
   792
    I.E. THE RESPECTIVE ASSUMPTION IS STORED AND Rule.Thm APPLIED; WN051028 
neuper@37950
   793
    ... WN0609???MG.*)
s1210629013@55444
   794
val rat_mult_div_pow = prep_rls'(
walther@59857
   795
  Rule_Def.Repeat {id = "rat_mult_div_pow", preconds = [], rew_ord = ("dummy_ord",Rewrite_Ord.dummy_ord), 
walther@59852
   796
    erls = Rule_Set.empty, srls = Rule_Set.Empty, calc = [], errpatts = [],
walther@59871
   797
    rules = [Rule.Thm ("rat_mult", ThmC.numerals_to_Free @{thm rat_mult}),
neuper@52105
   798
      (*"?a / ?b * (?c / ?d) = ?a * ?c / (?b * ?d)"*)
walther@59871
   799
      Rule.Thm ("rat_mult_poly_l", ThmC.numerals_to_Free @{thm rat_mult_poly_l}),
neuper@52105
   800
      (*"?c is_polyexp ==> ?c * (?a / ?b) = ?c * ?a / ?b"*)
walther@59871
   801
      Rule.Thm ("rat_mult_poly_r", ThmC.numerals_to_Free @{thm rat_mult_poly_r}),
neuper@52105
   802
      (*"?c is_polyexp ==> ?a / ?b * ?c = ?a * ?c / ?b"*)
neuper@52105
   803
      
walther@59871
   804
      Rule.Thm ("real_divide_divide1_mg", ThmC.numerals_to_Free @{thm real_divide_divide1_mg}),
neuper@52105
   805
      (*"y ~= 0 ==> (u / v) / (y / z) = (u * z) / (y * v)"*)
walther@59871
   806
      Rule.Thm ("divide_divide_eq_right", ThmC.numerals_to_Free @{thm divide_divide_eq_right}),
neuper@52105
   807
      (*"?x / (?y / ?z) = ?x * ?z / ?y"*)
walther@59871
   808
      Rule.Thm ("divide_divide_eq_left", ThmC.numerals_to_Free @{thm divide_divide_eq_left}),
neuper@52105
   809
      (*"?x / ?y / ?z = ?x / (?y * ?z)"*)
walther@59878
   810
      Rule.Eval ("Rings.divide_class.divide", Prog_Expr.eval_cancel "#divide_e"),
neuper@52105
   811
      
walther@59871
   812
      Rule.Thm ("rat_power", ThmC.numerals_to_Free @{thm rat_power})
walther@60260
   813
      (*"(?a / ?b)  \<up>  ?n = ?a  \<up>  ?n / ?b  \<up>  ?n"*)
neuper@52105
   814
      ],
walther@59878
   815
    scr = Rule.Empty_Prog});
neuper@37950
   816
s1210629013@55444
   817
val rat_reduce_1 = prep_rls'(
walther@59857
   818
  Rule_Def.Repeat {id = "rat_reduce_1", preconds = [], rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord), 
walther@59852
   819
    erls = Rule_Set.empty, srls = Rule_Set.Empty, calc = [], errpatts = [], 
neuper@52105
   820
    rules = 
walther@59871
   821
      [Rule.Thm ("div_by_1", ThmC.numerals_to_Free @{thm div_by_1}),
neuper@52105
   822
      (*"?x / 1 = ?x"*)
walther@59871
   823
      Rule.Thm ("mult_1_left", ThmC.numerals_to_Free @{thm mult_1_left})           
neuper@52105
   824
      (*"1 * z = z"*)
neuper@52105
   825
      ],
walther@59878
   826
    scr = Rule.Empty_Prog});
neuper@52105
   827
neuper@52105
   828
(* looping part of norm_Rational *)
s1210629013@55444
   829
val norm_Rational_rls = prep_rls' (
walther@59857
   830
  Rule_Def.Repeat {id = "norm_Rational_rls", preconds = [], rew_ord = ("dummy_ord",Rewrite_Ord.dummy_ord), 
walther@59851
   831
    erls = norm_rat_erls, srls = Rule_Set.Empty, calc = [], errpatts = [],
wneuper@59416
   832
    rules = [Rule.Rls_ add_fractions_p_rls,
wneuper@59416
   833
      Rule.Rls_ rat_mult_div_pow,
wneuper@59416
   834
      Rule.Rls_ make_rat_poly_with_parentheses,
wneuper@59416
   835
      Rule.Rls_ cancel_p_rls,
wneuper@59416
   836
      Rule.Rls_ rat_reduce_1
neuper@52105
   837
      ],
walther@59878
   838
    scr = Rule.Empty_Prog});
neuper@52105
   839
s1210629013@55444
   840
val norm_Rational = prep_rls' (
walther@59878
   841
  Rule_Set.Sequence 
walther@59857
   842
    {id = "norm_Rational", preconds = [], rew_ord = ("dummy_ord", Rewrite_Ord.dummy_ord), 
walther@59851
   843
    erls = norm_rat_erls, srls = Rule_Set.Empty, calc = [], errpatts = [],
wneuper@59416
   844
    rules = [Rule.Rls_ discard_minus,
wneuper@59416
   845
      Rule.Rls_ rat_mult_poly,             (* removes double fractions like a/b/c *)
wneuper@59416
   846
      Rule.Rls_ make_rat_poly_with_parentheses,
wneuper@59416
   847
      Rule.Rls_ cancel_p_rls,
wneuper@59416
   848
      Rule.Rls_ norm_Rational_rls,         (* the main rls, looping (#) *)
wneuper@59416
   849
      Rule.Rls_ discard_parentheses1       (* mult only *)
neuper@52100
   850
      ],
walther@59878
   851
    scr = Rule.Empty_Prog});
wneuper@59472
   852
\<close>
neuper@52125
   853
wneuper@59472
   854
setup \<open>KEStore_Elems.add_rlss 
neuper@52125
   855
  [("calculate_Rational", (Context.theory_name @{theory}, calculate_Rational)), 
neuper@52125
   856
  ("calc_rat_erls", (Context.theory_name @{theory}, calc_rat_erls)), 
neuper@52125
   857
  ("rational_erls", (Context.theory_name @{theory}, rational_erls)), 
neuper@52125
   858
  ("cancel_p", (Context.theory_name @{theory}, cancel_p)), 
neuper@52125
   859
  ("add_fractions_p", (Context.theory_name @{theory}, add_fractions_p)),
neuper@52125
   860
 
neuper@52125
   861
  ("add_fractions_p_rls", (Context.theory_name @{theory}, add_fractions_p_rls)), 
neuper@52125
   862
  ("powers_erls", (Context.theory_name @{theory}, powers_erls)), 
neuper@52125
   863
  ("powers", (Context.theory_name @{theory}, powers)), 
neuper@52125
   864
  ("rat_mult_divide", (Context.theory_name @{theory}, rat_mult_divide)), 
neuper@52125
   865
  ("reduce_0_1_2", (Context.theory_name @{theory}, reduce_0_1_2)),
neuper@52125
   866
 
neuper@52125
   867
  ("rat_reduce_1", (Context.theory_name @{theory}, rat_reduce_1)), 
neuper@52125
   868
  ("norm_rat_erls", (Context.theory_name @{theory}, norm_rat_erls)), 
neuper@52125
   869
  ("norm_Rational", (Context.theory_name @{theory}, norm_Rational)), 
neuper@52125
   870
  ("norm_Rational_rls", (Context.theory_name @{theory}, norm_Rational_rls)), 
neuper@55493
   871
  ("norm_Rational_min", (Context.theory_name @{theory}, norm_Rational_min)),
neuper@52125
   872
  ("norm_Rational_parenthesized", (Context.theory_name @{theory}, norm_Rational_parenthesized)),
neuper@52125
   873
 
neuper@52125
   874
  ("rat_mult_poly", (Context.theory_name @{theory}, rat_mult_poly)), 
neuper@52125
   875
  ("rat_mult_div_pow", (Context.theory_name @{theory}, rat_mult_div_pow)), 
wneuper@59472
   876
  ("cancel_p_rls", (Context.theory_name @{theory}, cancel_p_rls))]\<close>
neuper@37950
   877
wneuper@59472
   878
section \<open>A problem for simplification of rationals\<close>
wneuper@59472
   879
setup \<open>KEStore_Elems.add_pbts
walther@59973
   880
  [(Problem.prep_input thy "pbl_simp_rat" [] Problem.id_empty
walther@59997
   881
      (["rational", "simplification"],
s1210629013@55339
   882
        [("#Given" ,["Term t_t"]),
s1210629013@55339
   883
          ("#Where" ,["t_t is_ratpolyexp"]),
s1210629013@55339
   884
          ("#Find"  ,["normalform n_n"])],
walther@59852
   885
        Rule_Set.append_rules "empty" Rule_Set.empty [(*for preds in where_*)], 
walther@59997
   886
        SOME "Simplify t_t", [["simplification", "of_rationals"]]))]\<close>
neuper@37950
   887
wneuper@59472
   888
section \<open>A methods for simplification of rationals\<close>
s1210629013@55373
   889
(*WN061025 this methods script is copied from (auto-generated) script
s1210629013@55373
   890
  of norm_Rational in order to ease repair on inform*)
wneuper@59545
   891
wneuper@59504
   892
partial_function (tailrec) simplify :: "real \<Rightarrow> real"
wneuper@59504
   893
  where
walther@59716
   894
"simplify term = (
walther@59637
   895
  (Try (Rewrite_Set ''discard_minus'') #>
walther@59637
   896
   Try (Rewrite_Set ''rat_mult_poly'') #>
walther@59637
   897
   Try (Rewrite_Set ''make_rat_poly_with_parentheses'') #>
walther@59637
   898
   Try (Rewrite_Set ''cancel_p_rls'') #>
walther@59635
   899
   (Repeat (
walther@59637
   900
     Try (Rewrite_Set ''add_fractions_p_rls'') #>
walther@59637
   901
     Try (Rewrite_Set ''rat_mult_div_pow'') #>
walther@59637
   902
     Try (Rewrite_Set ''make_rat_poly_with_parentheses'') #>
walther@59637
   903
     Try (Rewrite_Set ''cancel_p_rls'') #>
walther@59637
   904
     Try (Rewrite_Set ''rat_reduce_1''))) #>
walther@59635
   905
   Try (Rewrite_Set ''discard_parentheses1''))
walther@59716
   906
   term)"
wneuper@59472
   907
setup \<open>KEStore_Elems.add_mets
walther@60154
   908
    [MethodC.prep_input thy "met_simp_rat" [] MethodC.id_empty
walther@59997
   909
      (["simplification", "of_rationals"],
s1210629013@55373
   910
        [("#Given" ,["Term t_t"]),
s1210629013@55373
   911
          ("#Where" ,["t_t is_ratpolyexp"]),
s1210629013@55373
   912
          ("#Find"  ,["normalform n_n"])],
walther@59852
   913
	      {rew_ord'="tless_true", rls' = Rule_Set.empty, calc = [], srls = Rule_Set.empty, 
walther@59852
   914
	        prls = Rule_Set.append_rules "simplification_of_rationals_prls" Rule_Set.empty 
walther@60278
   915
				    [(*for preds in where_*) Rule.Eval ("Rational.is_ratpolyexp", eval_is_ratpolyexp "")],
walther@59852
   916
				  crls = Rule_Set.empty, errpats = [], nrls = norm_Rational_rls},
wneuper@59551
   917
				  @{thm simplify.simps})]
walther@60278
   918
\<close> ML \<open>
walther@60278
   919
\<close> ML \<open>
wneuper@59472
   920
\<close>
neuper@52105
   921
end