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