src/Tools/isac/Knowledge/Atools.thy
author Walther Neuper <neuper@ist.tugraz.at>
Sun, 14 Oct 2012 21:26:02 +0200
changeset 48764 fd9145fbe471
parent 48763 9b9936d79dbe
child 48789 498ed5bb1004
permissions -rw-r--r--
2011-->2012: "isabelle usedir -b HOL Isac" works
neuper@37944
     1
(* Title:  tools for arithmetic
neuper@37944
     2
   Author: Walther Neuper 010308
neuper@37944
     3
   (c) due to copyright terms
neuper@37906
     4
neuper@37948
     5
12345678901234567890123456789012345678901234567890123456789012345678901234567890
neuper@37948
     6
        10        20        30        40        50        60        70        80
neuper@37906
     7
*)
neuper@37906
     8
bonzai@41919
     9
theory Atools imports Descript "../Interpret/Interpret" "../xmlsrc/xmlsrc"
neuper@37967
    10
begin
neuper@37906
    11
neuper@37906
    12
consts
neuper@37906
    13
neuper@37946
    14
  Arbfix           :: "real"
neuper@37949
    15
  Undef            :: "real"
neuper@37949
    16
  dummy            :: "real"
neuper@37949
    17
neuper@37949
    18
  some'_occur'_in  :: "[real list, 'a] => bool" ("some'_of _ occur'_in _")
neuper@37949
    19
  occurs'_in       :: "[real     , 'a] => bool" ("_ occurs'_in _")
neuper@37949
    20
neuper@37949
    21
  pow              :: "[real, real] => real"    (infixr "^^^" 80)
neuper@37949
    22
(* ~~~ power doesn't allow Free("2",real) ^ Free("2",nat)
neuper@37949
    23
                           ~~~~     ~~~~    ~~~~     ~~~*)
neuper@37949
    24
(*WN0603 at FE-interface encoded strings to '^', 
neuper@37949
    25
	see 'fun encode', fun 'decode'*)
neuper@37949
    26
neuper@37949
    27
  abs              :: "real => real"            ("(|| _ ||)")
neuper@37949
    28
(* ~~~ FIXXXME Isabelle2002 has abs already !!!*)
neuper@37949
    29
  absset           :: "real set => real"        ("(||| _ |||)")
neuper@37949
    30
  (*is numeral constant ?*)
neuper@37949
    31
  is'_const        :: "real => bool"            ("_ is'_const" 10)
neuper@37949
    32
  (*is_const rename to is_num FIXXXME.WN.16.5.03 *)
neuper@37949
    33
  is'_atom         :: "real => bool"            ("_ is'_atom" 10)
neuper@37949
    34
  is'_even         :: "real => bool"            ("_ is'_even" 10)
neuper@37949
    35
		
neuper@37949
    36
  (* identity on term level*)
neuper@37949
    37
  ident            :: "['a, 'a] => bool"        ("(_ =!=/ _)" [51, 51] 50)
neuper@37949
    38
neuper@37949
    39
  argument'_in     :: "real => real"            ("argument'_in _" 10)
neuper@37949
    40
  sameFunId        :: "[real, bool] => bool"    (**"same'_funid _ _" 10
neuper@37949
    41
	WN0609 changed the id, because ".. _ _" inhibits currying**)
neuper@37949
    42
  filter'_sameFunId:: "[real, bool list] => bool list" 
neuper@37949
    43
					        ("filter'_sameFunId _ _" 10)
neuper@37949
    44
  boollist2sum     :: "bool list => real"
neuper@37949
    45
t@42211
    46
axioms(*axiomatization where*) (*for evaluating the assumptions of conditional rules*)
neuper@37949
    47
t@42211
    48
  last_thmI:	      "lastI (x#xs) = (if xs =!= [] then x else lastI xs)" (*and*)
t@42211
    49
  real_unari_minus:   "- a = (-1) * a" (*and*)
neuper@37949
    50
t@42211
    51
  rle_refl:           "(n::real) <= n" (*and*)
t@42211
    52
  radd_left_cancel_le:"((k::real) + m <= k + n) = (m <= n)" (*and*)
t@42211
    53
  not_true:           "(~ True) = False" (*and*)
neuper@37965
    54
  not_false:          "(~ False) = True"
t@42211
    55
axioms(*axiomatization where*) (*..if replaced by "and" we get error:
t@42203
    56
Type unification failed: No type arity bool :: zero ...*)
t@42211
    57
  and_true:           "(a & True) = a" (*and*)
t@42211
    58
  and_false:          "(a & False) = False" (*and*)
t@42211
    59
  or_true:            "(a | True) = True" (*and*)
t@42211
    60
  or_false:           "(a | False) = a" (*and*)
t@42211
    61
  and_commute:        "(a & b) = (b & a)" (*and*)
t@42203
    62
  or_commute:         "(a | b) = (b | a)" 
neuper@37949
    63
neuper@37949
    64
  (*.should be in Rational.thy, but: 
neuper@37949
    65
   needed for asms in e.g. d2_pqformula1 in PolyEq.ML, RootEq.ML.*)
t@42211
    66
axioms(*axiomatization where*) (*..if replaced by "and" we get error:
t@42203
    67
Type unification failed: No type arity bool :: zero ...*)
neuper@37965
    68
  rat_leq1:	      "[| b ~= 0; d ~= 0 |] ==>
t@42211
    69
		       ((a / b) <= (c / d)) = ((a*d) <= (b*c))"(*Isa?*) (*and*)
neuper@37965
    70
  rat_leq2:	      "d ~= 0 ==>
t@42211
    71
		       ( a      <= (c / d)) = ((a*d) <=    c )"(*Isa?*) (*and*)
neuper@37965
    72
  rat_leq3:	      "b ~= 0 ==>
neuper@37949
    73
		       ((a / b) <=  c     ) = ( a    <= (b*c))"(*Isa?*)
neuper@37949
    74
neuper@37965
    75
neuper@37949
    76
text {*copy from doc/math-eng.tex WN.28.3.03
neuper@37949
    77
WN071228 extended *}
neuper@37949
    78
neuper@37949
    79
section {*Coding standards*}
neuper@37949
    80
subsection {*Identifiers*}
neuper@37949
    81
text {*Naming is particularily crucial, because Isabelles name space is global, and isac does not yet use the novel locale features introduces by Isar. For instance, {\tt probe} sounds reasonable as (1) a description in the model of a problem-pattern, (2) as an element of the problem hierarchies key, (3) as a socalled CAS-command, (4) as the name of a related script etc. However, all the cases (1)..(4) require different typing for one and the same identifier {\tt probe} which is impossible, and actually leads to strange errors (for instance (1) is used as string, except in a script addressing a Subproblem).
neuper@37949
    82
neuper@37949
    83
This are the preliminary rules for naming identifiers>
neuper@37949
    84
\begin{description}
neuper@37949
    85
\item [elements of a key] into the hierarchy of problems or methods must not contain capital letters and may contain underscrores, e.g. {\tt probe, for_polynomials}.
neuper@37949
    86
\item [descriptions in problem-patterns] must contain at least 1 capital letter and must not contain underscores, e.g. {\tt Probe, forPolynomials}.
neuper@37949
    87
\item [CAS-commands] follow the same rules as descriptions in problem-patterns above, thus beware of conflicts~!
neuper@37949
    88
\item [script identifiers] always end with {\tt Script}, e.g. {\tt ProbeScript}.
neuper@37949
    89
\item [???] ???
neuper@37949
    90
\item [???] ???
neuper@37949
    91
\end{description}
neuper@37949
    92
%WN071228 extended *}
neuper@37949
    93
neuper@37949
    94
subsection {*Rule sets*}
neuper@37949
    95
text {*The actual version of the coding standards for rulesets is in {\tt /IsacKnowledge/Atools.ML where it can be viewed using the knowledge browsers.
neuper@37949
    96
neuper@37949
    97
There are rulesets visible to the student, and there are rulesets visible (in general) only for math authors. There are also rulesets which {\em must} exist for {\em each} theory; these contain the identifier of the respective theory (including all capital letters) as indicated by {\it Thy} below.
neuper@37949
    98
\begin{description}
neuper@37949
    99
neuper@37949
   100
\item [norm\_{\it Thy}] exists for each theory, and {\em efficiently} calculates a normalform for all terms which can be expressed by the definitions of the respective theory (and the respective parents).
neuper@37949
   101
neuper@37949
   102
\item [simplify\_{\it Thy}] exists for each theory, and calculates a normalform for all terms which can be expressed by the definitions of the respective theory (and the respective parents) such, that the rewrites can be presented to the student.
neuper@37949
   103
neuper@37949
   104
\item [calculate\_{\it Thy}] exists for each theory, and evaluates terms with numerical constants only (i.e. all terms which can be expressed by the definitions of the respective theory and the respective parent theories). In particular, this ruleset includes evaluating in/equalities with numerical constants only.
neuper@37949
   105
WN.3.7.03: may be dropped due to more generality: numericals and non-numericals are logically equivalent, where the latter often add to the assumptions (e.g. in Check_elementwise).
neuper@37949
   106
\end{description}
neuper@37949
   107
neuper@37949
   108
The above rulesets are all visible to the user, and also may be input; thus they must be contained in the global associationlist {\tt ruleset':= }~! All these rulesets must undergo a preparation using the function {\tt prep_rls}, which generates a script for stepwise rewriting etc.
neuper@37949
   109
The following rulesets are used for internal purposes and usually invisible to the (naive) user:
neuper@37949
   110
\begin{description}
neuper@37949
   111
neuper@37949
   112
\item [*\_erls] TODO
neuper@37949
   113
\item [*\_prls] 
neuper@37949
   114
\item [*\_srls] 
neuper@37949
   115
neuper@37949
   116
\end{description}
neuper@37949
   117
{\tt append_rls, merge_rls, remove_rls} TODO
neuper@37949
   118
*}
neuper@37946
   119
neuper@37946
   120
ML {*
neuper@37972
   121
val thy = @{theory};
neuper@37949
   122
neuper@37949
   123
(** evaluation of numerals and special predicates on the meta-level **)
neuper@37949
   124
(*-------------------------functions---------------------*)
neuper@37949
   125
local (* rlang 09.02 *)
neuper@37949
   126
    (*.a 'c is coefficient of v' if v does occur in c.*)
neuper@37949
   127
    fun coeff_in v c = member op = (vars c) v;
neuper@37949
   128
in
neuper@37949
   129
    fun occurs_in v t = coeff_in v t;
neuper@37949
   130
end;
neuper@37949
   131
neuper@37949
   132
(*("occurs_in", ("Atools.occurs'_in", eval_occurs_in ""))*)
neuper@37949
   133
fun eval_occurs_in _ "Atools.occurs'_in"
neuper@37949
   134
	     (p as (Const ("Atools.occurs'_in",_) $ v $ t)) _ =
neuper@38015
   135
    ((*tracing("@@@ eval_occurs_in: v= "^(term2str v));
neuper@38015
   136
     tracing("@@@ eval_occurs_in: t= "^(term2str t));*)
neuper@37949
   137
     if occurs_in v t
neuper@37949
   138
    then SOME ((term2str p) ^ " = True",
neuper@48760
   139
	  Trueprop $ (mk_equality (p, @{term True})))
neuper@37949
   140
    else SOME ((term2str p) ^ " = False",
neuper@48760
   141
	  Trueprop $ (mk_equality (p, @{term False}))))
neuper@37949
   142
  | eval_occurs_in _ _ _ _ = NONE;
neuper@37949
   143
neuper@37949
   144
(*some of the (bound) variables (eg. in an eqsys) "vs" occur in term "t"*)   
neuper@37949
   145
fun some_occur_in vs t = 
neuper@37949
   146
    let fun occurs_in' a b = occurs_in b a
neuper@37949
   147
    in foldl or_ (false, map (occurs_in' t) vs) end;
neuper@37949
   148
neuper@37949
   149
(*("some_occur_in", ("Atools.some'_occur'_in", 
neuper@37949
   150
			eval_some_occur_in "#eval_some_occur_in_"))*)
neuper@37949
   151
fun eval_some_occur_in _ "Atools.some'_occur'_in"
neuper@37949
   152
			  (p as (Const ("Atools.some'_occur'_in",_) 
neuper@37949
   153
				       $ vs $ t)) _ =
neuper@37949
   154
    if some_occur_in (isalist2list vs) t
neuper@37949
   155
    then SOME ((term2str p) ^ " = True",
neuper@48760
   156
	       Trueprop $ (mk_equality (p, @{term True})))
neuper@37949
   157
    else SOME ((term2str p) ^ " = False",
neuper@48760
   158
	       Trueprop $ (mk_equality (p, @{term False})))
neuper@37949
   159
  | eval_some_occur_in _ _ _ _ = NONE;
neuper@37949
   160
neuper@37949
   161
neuper@37949
   162
neuper@37949
   163
neuper@37949
   164
(*evaluate 'is_atom'*)
neuper@37949
   165
(*("is_atom",("Atools.is'_atom",eval_is_atom "#is_atom_"))*)
neuper@37949
   166
fun eval_is_atom (thmid:string) "Atools.is'_atom"
neuper@37949
   167
		 (t as (Const(op0,_) $ arg)) thy = 
neuper@37949
   168
    (case arg of 
neuper@37949
   169
	 Free (n,_) => SOME (mk_thmid thmid op0 n "", 
neuper@48760
   170
			      Trueprop $ (mk_equality (t, @{term True})))
neuper@37949
   171
       | _ => SOME (mk_thmid thmid op0 "" "", 
neuper@48760
   172
		    Trueprop $ (mk_equality (t, @{term False}))))
neuper@37949
   173
  | eval_is_atom _ _ _ _ = NONE;
neuper@37949
   174
neuper@37949
   175
(*evaluate 'is_even'*)
neuper@37949
   176
fun even i = (i div 2) * 2 = i;
neuper@37949
   177
(*("is_even",("Atools.is'_even",eval_is_even "#is_even_"))*)
neuper@37949
   178
fun eval_is_even (thmid:string) "Atools.is'_even"
neuper@37949
   179
		 (t as (Const(op0,_) $ arg)) thy = 
neuper@37949
   180
    (case arg of 
neuper@37949
   181
	Free (n,_) =>
neuper@37949
   182
	 (case int_of_str n of
neuper@37949
   183
	      SOME i =>
neuper@37949
   184
	      if even i then SOME (mk_thmid thmid op0 n "", 
neuper@48760
   185
				   Trueprop $ (mk_equality (t, @{term True})))
neuper@37949
   186
	      else SOME (mk_thmid thmid op0 "" "", 
neuper@48760
   187
			 Trueprop $ (mk_equality (t, @{term False})))
neuper@37949
   188
	    | _ => NONE)
neuper@37949
   189
       | _ => NONE)
neuper@37949
   190
  | eval_is_even _ _ _ _ = NONE; 
neuper@37949
   191
neuper@37949
   192
(*evaluate 'is_const'*)
neuper@37949
   193
(*("is_const",("Atools.is'_const",eval_const "#is_const_"))*)
neuper@37949
   194
fun eval_const (thmid:string) _(*"Atools.is'_const" WN050820 diff.beh. rooteq*)
neuper@37949
   195
	       (t as (Const(op0,t0) $ arg)) (thy:theory) = 
neuper@37949
   196
    (*eval_const FIXXXXXME.WN.16.5.03 still forgets ComplexI*)
neuper@37949
   197
    (case arg of 
neuper@37949
   198
       Const (n1,_) =>
neuper@37949
   199
	 SOME (mk_thmid thmid op0 n1 "", 
neuper@48760
   200
	       Trueprop $ (mk_equality (t, @{term False})))
neuper@37949
   201
     | Free (n1,_) =>
neuper@37949
   202
	 if is_numeral n1
neuper@37949
   203
	   then SOME (mk_thmid thmid op0 n1 "", 
neuper@48760
   204
		      Trueprop $ (mk_equality (t, @{term True})))
neuper@37949
   205
	 else SOME (mk_thmid thmid op0 n1 "", 
neuper@48760
   206
		    Trueprop $ (mk_equality (t, @{term False})))
neuper@37949
   207
     | Const ("Float.Float",_) =>
neuper@37949
   208
       SOME (mk_thmid thmid op0 (term2str arg) "", 
neuper@48760
   209
	     Trueprop $ (mk_equality (t, @{term True})))
neuper@37949
   210
     | _ => (*NONE*)
neuper@37949
   211
       SOME (mk_thmid thmid op0 (term2str arg) "", 
neuper@48760
   212
		    Trueprop $ (mk_equality (t, @{term False}))))
neuper@37949
   213
  | eval_const _ _ _ _ = NONE; 
neuper@37949
   214
neuper@37949
   215
(*. evaluate binary, associative, commutative operators: *,+,^ .*)
neuper@38014
   216
(*("PLUS"    ,("Groups.plus_class.plus"        ,eval_binop "#add_")),
neuper@38034
   217
  ("TIMES"   ,("Groups.times_class.times"        ,eval_binop "#mult_")),
neuper@37949
   218
  ("POWER"  ,("Atools.pow"  ,eval_binop "#power_"))*)
neuper@37949
   219
neuper@37949
   220
(* val (thmid,op_,t as(Const (op0,t0) $ Free (n1,t1) $ Free(n2,t2)),thy) =
neuper@37949
   221
       ("xxxxxx",op_,t,thy);
neuper@37949
   222
   *)
neuper@37949
   223
fun mk_thmid_f thmid ((v11, v12), (p11, p12)) ((v21, v22), (p21, p22))  = 
neuper@37949
   224
    thmid ^ "Float ((" ^ 
neuper@37949
   225
    (string_of_int v11)^","^(string_of_int v12)^"), ("^
neuper@37949
   226
    (string_of_int p11)^","^(string_of_int p12)^")) __ (("^
neuper@37949
   227
    (string_of_int v21)^","^(string_of_int v22)^"), ("^
neuper@37949
   228
    (string_of_int p21)^","^(string_of_int p22)^"))";
neuper@37949
   229
neuper@37949
   230
(*.convert int and float to internal floatingpoint prepresentation.*)
neuper@37949
   231
fun numeral (Free (str, T)) = 
neuper@37949
   232
    (case int_of_str str of
neuper@37949
   233
	 SOME i => SOME ((i, 0), (0, 0))
neuper@37949
   234
       | NONE => NONE)
neuper@37949
   235
  | numeral (Const ("Float.Float", _) $
neuper@41972
   236
		   (Const ("Product_Type.Pair", _) $
neuper@41972
   237
			  (Const ("Product_Type.Pair", T) $ Free (v1, _) $ Free (v2,_)) $
neuper@41972
   238
			  (Const ("Product_Type.Pair", _) $ Free (p1, _) $ Free (p2,_))))=
neuper@37949
   239
    (case (int_of_str v1, int_of_str v2, int_of_str p1, int_of_str p2) of
neuper@37949
   240
	(SOME v1', SOME v2', SOME p1', SOME p2') =>
neuper@37949
   241
	SOME ((v1', v2'), (p1', p2'))
neuper@37949
   242
      | _ => NONE)
neuper@37949
   243
  | numeral _ = NONE;
neuper@37949
   244
neuper@37949
   245
(*.evaluate binary associative operations.*)
neuper@37949
   246
fun eval_binop (thmid:string) (op_:string) 
neuper@37949
   247
	       (t as ( Const(op0,t0) $ 
neuper@37949
   248
			    (Const(op0',t0') $ v $ t1) $ t2)) 
neuper@37949
   249
	       thy =                                     (*binary . (v.n1).n2*)
neuper@37949
   250
    if op0 = op0' then
neuper@37949
   251
	case (numeral t1, numeral t2) of
neuper@37949
   252
	    (SOME n1, SOME n2) =>
neuper@37949
   253
	    let val (T1,T2,Trange) = dest_binop_typ t0
neuper@38014
   254
		val res = calc (if op0 = "Groups.minus_class.minus" then "Groups.plus_class.plus" else op0) n1 n2
neuper@38014
   255
		(*WN071229 "Rings.inverse_class.divide" never tried*)
neuper@37949
   256
		val rhs = var_op_float v op_ t0 T1 res
neuper@37949
   257
		val prop = Trueprop $ (mk_equality (t, rhs))
neuper@37949
   258
	    in SOME (mk_thmid_f thmid n1 n2, prop) end
neuper@37949
   259
	  | _ => NONE
neuper@37949
   260
    else NONE
neuper@37949
   261
  | eval_binop (thmid:string) (op_:string) 
neuper@37949
   262
	       (t as 
neuper@37949
   263
		  (Const (op0, t0) $ t1 $ 
neuper@37949
   264
			 (Const (op0', t0') $ t2 $ v))) 
neuper@37949
   265
	       thy =                                     (*binary . n1.(n2.v)*)
neuper@37949
   266
  if op0 = op0' then
neuper@37949
   267
	case (numeral t1, numeral t2) of
neuper@37949
   268
	    (SOME n1, SOME n2) =>
neuper@38014
   269
	    if op0 = "Groups.minus_class.minus" then NONE else
neuper@37949
   270
	    let val (T1,T2,Trange) = dest_binop_typ t0
neuper@37949
   271
		val res = calc op0 n1 n2
neuper@37949
   272
		val rhs = float_op_var v op_ t0 T1 res
neuper@37949
   273
		val prop = Trueprop $ (mk_equality (t, rhs))
neuper@37949
   274
	    in SOME (mk_thmid_f thmid n1 n2, prop) end
neuper@37949
   275
	  | _ => NONE
neuper@37949
   276
  else NONE
neuper@37949
   277
    
neuper@37949
   278
  | eval_binop (thmid:string) (op_:string)
neuper@37949
   279
	       (t as (Const (op0,t0) $ t1 $ t2)) thy =       (*binary . n1.n2*)
neuper@37949
   280
    (case (numeral t1, numeral t2) of
neuper@37949
   281
	 (SOME n1, SOME n2) =>
neuper@37949
   282
	 let val (T1,T2,Trange) = dest_binop_typ t0;
neuper@37949
   283
	     val res = calc op0 n1 n2;
neuper@37949
   284
	     val rhs = term_of_float Trange res;
neuper@37949
   285
	     val prop = Trueprop $ (mk_equality (t, rhs));
neuper@37949
   286
	 in SOME (mk_thmid_f thmid n1 n2, prop) end
neuper@37949
   287
       | _ => NONE)
neuper@37949
   288
  | eval_binop _ _ _ _ = NONE; 
neuper@37949
   289
(*
neuper@38014
   290
> val SOME (thmid, t) = eval_binop "#add_" "Groups.plus_class.plus" (str2term "-1 + 2") thy;
neuper@37949
   291
> term2str t;
neuper@37949
   292
val it = "-1 + 2 = 1"
neuper@37949
   293
> val t = str2term "-1 * (-1 * a)";
neuper@38034
   294
> val SOME (thmid, t) = eval_binop "#mult_" "Groups.times_class.times" t thy;
neuper@37949
   295
> term2str t;
neuper@37949
   296
val it = "-1 * (-1 * a) = 1 * a"*)
neuper@37949
   297
neuper@37949
   298
neuper@37949
   299
neuper@37949
   300
(*.evaluate < and <= for numerals.*)
neuper@38045
   301
(*("le"      ,("Orderings.ord_class.less"        ,eval_equ "#less_")),
neuper@38045
   302
  ("leq"     ,("Orderings.ord_class.less_eq"       ,eval_equ "#less_equal_"))*)
neuper@37949
   303
fun eval_equ (thmid:string) (op_:string) (t as 
neuper@37949
   304
	       (Const (op0,t0) $ Free (n1,t1) $ Free(n2,t2))) thy = 
neuper@37949
   305
    (case (int_of_str n1, int_of_str n2) of
neuper@37949
   306
	 (SOME n1', SOME n2') =>
neuper@37949
   307
  if calc_equ (strip_thy op0) (n1', n2')
neuper@37949
   308
    then SOME (mk_thmid thmid op0 n1 n2, 
neuper@48760
   309
	  Trueprop $ (mk_equality (t, @{term True})))
neuper@37949
   310
  else SOME (mk_thmid thmid op0 n1 n2,  
neuper@48760
   311
	  Trueprop $ (mk_equality (t, @{term False})))
neuper@37949
   312
       | _ => NONE)
neuper@37949
   313
    
neuper@37949
   314
  | eval_equ _ _ _ _ = NONE;
neuper@37949
   315
neuper@37949
   316
neuper@37949
   317
(*evaluate identity
neuper@37949
   318
> reflI;
neuper@37949
   319
val it = "(?t = ?t) = True"
neuper@37949
   320
> val t = str2term "x = 0";
neuper@37949
   321
> val NONE = rewrite_ thy dummy_ord e_rls false reflI t;
neuper@37949
   322
neuper@37949
   323
> val t = str2term "1 = 0";
neuper@37949
   324
> val NONE = rewrite_ thy dummy_ord e_rls false reflI t;
neuper@37949
   325
----------- thus needs Calc !
neuper@37949
   326
> val t = str2term "0 = 0";
neuper@37949
   327
> val SOME (t',_) = rewrite_ thy dummy_ord e_rls false reflI t;
neuper@37949
   328
> term2str t';
neuper@41928
   329
val it = "HOL.True"
neuper@37949
   330
neuper@37949
   331
val t = str2term "Not (x = 0)";
neuper@37949
   332
atomt t; term2str t;
neuper@37949
   333
*** -------------
neuper@37949
   334
*** Const ( Not)
neuper@37949
   335
*** . Const ( op =)
neuper@37949
   336
*** . . Free ( x, )
neuper@37949
   337
*** . . Free ( 0, )
neuper@37949
   338
val it = "x ~= 0" : string*)
neuper@37949
   339
neuper@37949
   340
(*.evaluate identity on the term-level, =!= ,i.e. without evaluation of 
neuper@37949
   341
  the arguments: thus special handling by 'fun eval_binop'*)
neuper@37949
   342
(*("ident"   ,("Atools.ident",eval_ident "#ident_")):calc*)
neuper@37949
   343
fun eval_ident (thmid:string) "Atools.ident" (t as 
neuper@37949
   344
	       (Const (op0,t0) $ t1 $ t2 )) thy = 
neuper@37949
   345
  if t1 = t2
neuper@38053
   346
  then SOME (mk_thmid thmid op0 
neuper@38053
   347
	              ("(" ^ (Print_Mode.setmp [] (Syntax.string_of_term
neuper@38053
   348
                                                   (thy2ctxt thy)) t1) ^ ")")
neuper@38053
   349
	              ("(" ^ (Print_Mode.setmp [] (Syntax.string_of_term
neuper@38053
   350
                                                   (thy2ctxt thy)) t2) ^ ")"), 
neuper@48760
   351
	     Trueprop $ (mk_equality (t, @{term True})))
neuper@37949
   352
  else SOME (mk_thmid thmid op0  
neuper@38053
   353
	              ("(" ^ (Print_Mode.setmp [] (Syntax.string_of_term
neuper@38053
   354
                                                   (thy2ctxt thy)) t1) ^ ")")
neuper@38053
   355
	              ("(" ^ (Print_Mode.setmp [] (Syntax.string_of_term
neuper@38053
   356
                                                   (thy2ctxt thy)) t2) ^ ")"),  
neuper@48760
   357
	     Trueprop $ (mk_equality (t, @{term False})))
neuper@37949
   358
  | eval_ident _ _ _ _ = NONE;
neuper@37949
   359
(* TODO
neuper@37949
   360
> val t = str2term "x =!= 0";
neuper@37949
   361
> val SOME (str, t') = eval_ident "ident_" "b" t thy;
neuper@37949
   362
> term2str t';
neuper@37949
   363
val str = "ident_(x)_(0)" : string
neuper@37949
   364
val it = "(x =!= 0) = False" : string                                
neuper@37949
   365
> val t = str2term "1 =!= 0";
neuper@37949
   366
> val SOME (str, t') = eval_ident "ident_" "b" t thy;
neuper@37949
   367
> term2str t';
neuper@37949
   368
val str = "ident_(1)_(0)" : string 
neuper@37949
   369
val it = "(1 =!= 0) = False" : string                                       
neuper@37949
   370
> val t = str2term "0 =!= 0";
neuper@37949
   371
> val SOME (str, t') = eval_ident "ident_" "b" t thy;
neuper@37949
   372
> term2str t';
neuper@37949
   373
val str = "ident_(0)_(0)" : string
neuper@37949
   374
val it = "(0 =!= 0) = True" : string
neuper@37949
   375
*)
neuper@37949
   376
neuper@37949
   377
(*.evaluate identity of terms, which stay ready for evaluation in turn;
neuper@37949
   378
  thus returns False only for atoms.*)
neuper@41922
   379
(*("equal"   ,("HOL.eq",eval_equal "#equal_")):calc*)
neuper@41922
   380
fun eval_equal (thmid : string) "HOL.eq" (t as 
neuper@37949
   381
	       (Const (op0,t0) $ t1 $ t2 )) thy = 
neuper@38053
   382
    if t1 = t2
neuper@38053
   383
    then SOME (mk_thmid thmid op0 
neuper@38053
   384
	                ("(" ^ Print_Mode.setmp [] (Syntax.string_of_term
neuper@38053
   385
                                                        (thy2ctxt thy)) t1 ^
neuper@38053
   386
                         ")")
neuper@38053
   387
	                ("(" ^ Print_Mode.setmp [] (Syntax.string_of_term
neuper@38053
   388
                                                        (thy2ctxt thy)) t2 ^
neuper@38053
   389
                         ")"), 
neuper@48760
   390
	       Trueprop $ (mk_equality (t, @{term True})))
neuper@38053
   391
    else (case (is_atom t1, is_atom t2) of
neuper@38053
   392
	      (true, true) => 
neuper@38053
   393
	      SOME (mk_thmid thmid op0  
neuper@38053
   394
			     ("(" ^ Print_Mode.setmp [] (Syntax.string_of_term
neuper@38053
   395
                                                             (thy2ctxt thy)) t1^
neuper@38053
   396
                              ")") ("(" ^ 
neuper@38053
   397
                                    Print_Mode.setmp [] (Syntax.string_of_term
neuper@38053
   398
                                                             (thy2ctxt thy)) t2^
neuper@38053
   399
                                    ")"),
neuper@48760
   400
		    Trueprop $ (mk_equality (t, @{term False})))
neuper@38053
   401
	    | _ => NONE)                     (* NOT is_atom t1,t2 --> rew_sub *)
neuper@38053
   402
  | eval_equal _ _ _ _ = NONE;                                  (* error-exit *)
neuper@37949
   403
(*
neuper@37949
   404
val t = str2term "x ~= 0";
neuper@37949
   405
val NONE = eval_equal "equal_" "b" t thy;
neuper@37949
   406
neuper@37949
   407
neuper@37949
   408
> val t = str2term "(x + 1) = (x + 1)";
neuper@37949
   409
> val SOME (str, t') = eval_equal "equal_" "b" t thy;
neuper@37949
   410
> term2str t';
neuper@37949
   411
val str = "equal_(x + 1)_(x + 1)" : string
neuper@37949
   412
val it = "(x + 1 = x + 1) = True" : string
neuper@37949
   413
> val t = str2term "x = 0";
neuper@37949
   414
> val NONE = eval_equal "equal_" "b" t thy;
neuper@37949
   415
neuper@37949
   416
> val t = str2term "1 = 0";
neuper@37949
   417
> val SOME (str, t') = eval_equal "equal_" "b" t thy;
neuper@37949
   418
> term2str t';
neuper@37949
   419
val str = "equal_(1)_(0)" : string 
neuper@37949
   420
val it = "(1 = 0) = False" : string
neuper@37949
   421
> val t = str2term "0 = 0";
neuper@37949
   422
> val SOME (str, t') = eval_equal "equal_" "b" t thy;
neuper@37949
   423
> term2str t';
neuper@37949
   424
val str = "equal_(0)_(0)" : string
neuper@37949
   425
val it = "(0 = 0) = True" : string
neuper@37949
   426
*)
neuper@37949
   427
neuper@37949
   428
neuper@42451
   429
*}
neuper@42451
   430
neuper@42451
   431
ML {*
neuper@37949
   432
(** evaluation on the metalevel **)
neuper@37949
   433
neuper@37949
   434
(*. evaluate HOL.divide .*)
neuper@38014
   435
(*("DIVIDE" ,("Rings.inverse_class.divide"  ,eval_cancel "#divide_e"))*)
neuper@38014
   436
fun eval_cancel (thmid:string) "Rings.inverse_class.divide" (t as 
neuper@37949
   437
	       (Const (op0,t0) $ Free (n1,t1) $ Free(n2,t2))) thy = 
neuper@37949
   438
    (case (int_of_str n1, int_of_str n2) of
neuper@37949
   439
	 (SOME n1', SOME n2') =>
neuper@37949
   440
  let 
neuper@37949
   441
    val sg = sign2 n1' n2';
neuper@37949
   442
    val (T1,T2,Trange) = dest_binop_typ t0;
neuper@37949
   443
    val gcd' = gcd (abs n1') (abs n2');
neuper@37949
   444
  in if gcd' = abs n2' 
neuper@37949
   445
     then let val rhs = term_of_num Trange (sg * (abs n1') div gcd')
neuper@37949
   446
	      val prop = Trueprop $ (mk_equality (t, rhs))
neuper@37949
   447
	  in SOME (mk_thmid thmid op0 n1 n2, prop) end     
neuper@37949
   448
     else if 0 < n2' andalso gcd' = 1 then NONE
neuper@37949
   449
     else let val rhs = num_op_num T1 T2 (op0,t0) (sg * (abs n1') div gcd')
neuper@37949
   450
				   ((abs n2') div gcd')
neuper@37949
   451
	      val prop = Trueprop $ (mk_equality (t, rhs))
neuper@37949
   452
	  in SOME (mk_thmid thmid op0 n1 n2, prop) end
neuper@37949
   453
  end
neuper@38015
   454
       | _ => ((*tracing"@@@ eval_cancel NONE";*)NONE))
neuper@37949
   455
neuper@37949
   456
  | eval_cancel _ _ _ _ = NONE;
neuper@37949
   457
neuper@37949
   458
(*. get the argument from a function-definition.*)
neuper@37949
   459
(*("argument_in" ,("Atools.argument'_in",
neuper@37949
   460
		   eval_argument_in "Atools.argument'_in"))*)
neuper@37949
   461
fun eval_argument_in _ "Atools.argument'_in" 
neuper@37949
   462
		     (t as (Const ("Atools.argument'_in", _) $ (f $ arg))) _ =
neuper@37949
   463
    if is_Free arg (*could be something to be simplified before*)
neuper@37949
   464
    then SOME (term2str t ^ " = " ^ term2str arg,
neuper@37949
   465
	       Trueprop $ (mk_equality (t, arg)))
neuper@37949
   466
    else NONE
neuper@37949
   467
  | eval_argument_in _ _ _ _ = NONE;
neuper@37949
   468
neuper@37949
   469
(*.check if the function-identifier of the first argument matches 
neuper@37949
   470
   the function-identifier of the lhs of the second argument.*)
neuper@37949
   471
(*("sameFunId" ,("Atools.sameFunId",
neuper@37949
   472
		   eval_same_funid "Atools.sameFunId"))*)
neuper@37949
   473
fun eval_sameFunId _ "Atools.sameFunId" 
neuper@37949
   474
		     (p as Const ("Atools.sameFunId",_) $ 
neuper@37949
   475
			(f1 $ _) $ 
neuper@41922
   476
			(Const ("HOL.eq", _) $ (f2 $ _) $ _)) _ =
neuper@37949
   477
    if f1 = f2 
neuper@37949
   478
    then SOME ((term2str p) ^ " = True",
neuper@48760
   479
	       Trueprop $ (mk_equality (p, @{term True})))
neuper@37949
   480
    else SOME ((term2str p) ^ " = False",
neuper@48760
   481
	       Trueprop $ (mk_equality (p, @{term False})))
neuper@37949
   482
| eval_sameFunId _ _ _ _ = NONE;
neuper@37949
   483
neuper@37949
   484
neuper@37949
   485
(*.from a list of fun-definitions "f x = ..." as 2nd argument
neuper@37949
   486
   filter the elements with the same fun-identfier in "f y"
neuper@37949
   487
   as the fst argument;
neuper@37949
   488
   this is, because Isabelles filter takes more than 1 sec.*)
neuper@41922
   489
fun same_funid f1 (Const ("HOL.eq", _) $ (f2 $ _) $ _) = f1 = f2
neuper@38031
   490
  | same_funid f1 t = error ("same_funid called with t = ("
neuper@37949
   491
				   ^term2str f1^") ("^term2str t^")");
neuper@37949
   492
(*("filter_sameFunId" ,("Atools.filter'_sameFunId",
neuper@37949
   493
		   eval_filter_sameFunId "Atools.filter'_sameFunId"))*)
neuper@37949
   494
fun eval_filter_sameFunId _ "Atools.filter'_sameFunId" 
neuper@37949
   495
		     (p as Const ("Atools.filter'_sameFunId",_) $ 
neuper@37949
   496
			(fid $ _) $ fs) _ =
neuper@37949
   497
    let val fs' = ((list2isalist HOLogic.boolT) o 
neuper@37949
   498
		   (filter (same_funid fid))) (isalist2list fs)
neuper@37949
   499
    in SOME (term2str (mk_equality (p, fs')),
neuper@37949
   500
	       Trueprop $ (mk_equality (p, fs'))) end
neuper@37949
   501
| eval_filter_sameFunId _ _ _ _ = NONE;
neuper@37949
   502
neuper@37949
   503
neuper@37949
   504
(*make a list of terms to a sum*)
neuper@37949
   505
fun list2sum [] = error ("list2sum called with []")
neuper@37949
   506
  | list2sum [s] = s
neuper@37949
   507
  | list2sum (s::ss) = 
neuper@37949
   508
    let fun sum su [s'] = 
neuper@38014
   509
	    Const ("Groups.plus_class.plus", [HOLogic.realT,HOLogic.realT]--->HOLogic.realT)
neuper@37949
   510
		  $ su $ s'
neuper@37949
   511
	  | sum su (s'::ss') = 
neuper@38014
   512
	    sum (Const ("Groups.plus_class.plus", [HOLogic.realT,HOLogic.realT]--->HOLogic.realT)
neuper@37949
   513
		  $ su $ s') ss'
neuper@37949
   514
    in sum s ss end;
neuper@37949
   515
neuper@37949
   516
(*make a list of equalities to the sum of the lhs*)
neuper@37949
   517
(*("boollist2sum"    ,("Atools.boollist2sum"    ,eval_boollist2sum "")):calc*)
neuper@37949
   518
fun eval_boollist2sum _ "Atools.boollist2sum" 
neuper@37949
   519
		      (p as Const ("Atools.boollist2sum", _) $ 
neuper@37949
   520
			 (l as Const ("List.list.Cons", _) $ _ $ _)) _ =
neuper@37949
   521
    let val isal = isalist2list l
neuper@37949
   522
	val lhss = map lhs isal
neuper@37949
   523
	val sum = list2sum lhss
neuper@37949
   524
    in SOME ((term2str p) ^ " = " ^ (term2str sum),
neuper@37949
   525
	  Trueprop $ (mk_equality (p, sum)))
neuper@37949
   526
    end
neuper@37949
   527
| eval_boollist2sum _ _ _ _ = NONE;
neuper@37949
   528
neuper@37949
   529
neuper@37949
   530
neuper@37949
   531
local
neuper@37949
   532
neuper@37949
   533
open Term;
neuper@37949
   534
neuper@37949
   535
in
neuper@37965
   536
fun termlessI (_:subst) uv = Term_Ord.termless uv;
neuper@37965
   537
fun term_ordI (_:subst) uv = Term_Ord.term_ord uv;
neuper@37949
   538
end;
neuper@37949
   539
neuper@37949
   540
neuper@37949
   541
(** rule set, for evaluating list-expressions in scripts 8.01.02 **)
neuper@37949
   542
neuper@37949
   543
neuper@37949
   544
val list_rls = 
neuper@37949
   545
    append_rls "list_rls" list_rls
neuper@38034
   546
	       [Calc ("Groups.times_class.times",eval_binop "#mult_"),
neuper@38014
   547
		Calc ("Groups.plus_class.plus", eval_binop "#add_"), 
neuper@38045
   548
		Calc ("Orderings.ord_class.less",eval_equ "#less_"),
neuper@38045
   549
		Calc ("Orderings.ord_class.less_eq",eval_equ "#less_equal_"),
neuper@37949
   550
		Calc ("Atools.ident",eval_ident "#ident_"),
neuper@41922
   551
		Calc ("HOL.eq",eval_equal "#equal_"),(*atom <> atom -> False*)
neuper@37949
   552
       
neuper@37949
   553
		Calc ("Tools.Vars",eval_var "#Vars_"),
neuper@37949
   554
		
neuper@37965
   555
		Thm ("if_True",num_str @{thm if_True}),
neuper@37965
   556
		Thm ("if_False",num_str @{thm if_False})
neuper@37949
   557
		];
neuper@37949
   558
neuper@37972
   559
ruleset' := overwritelthy thy (!ruleset',
neuper@37949
   560
  [("list_rls",list_rls)
neuper@37949
   561
   ]);
neuper@37949
   562
neuper@37949
   563
(*TODO.WN0509 reduce ids: tless_true = e_rew_ord' = e_rew_ord = dummy_ord*)
neuper@37949
   564
val tless_true = dummy_ord;
neuper@37949
   565
rew_ord' := overwritel (!rew_ord',
neuper@37949
   566
			[("tless_true", tless_true),
neuper@37949
   567
			 ("e_rew_ord'", tless_true),
neuper@37949
   568
			 ("dummy_ord", dummy_ord)]);
neuper@37949
   569
neuper@37949
   570
val calculate_Atools = 
neuper@37949
   571
    append_rls "calculate_Atools" e_rls
neuper@38045
   572
               [Calc ("Orderings.ord_class.less",eval_equ "#less_"),
neuper@38045
   573
		Calc ("Orderings.ord_class.less_eq",eval_equ "#less_equal_"),
neuper@41922
   574
		Calc ("HOL.eq",eval_equal "#equal_"),
neuper@37949
   575
neuper@37965
   576
		Thm  ("real_unari_minus",num_str @{thm real_unari_minus}),
neuper@38014
   577
		Calc ("Groups.plus_class.plus",eval_binop "#add_"),
neuper@38014
   578
		Calc ("Groups.minus_class.minus",eval_binop "#sub_"),
neuper@38034
   579
		Calc ("Groups.times_class.times",eval_binop "#mult_")
neuper@37949
   580
		];
neuper@37949
   581
neuper@42451
   582
*}
neuper@42451
   583
neuper@48763
   584
find_theorems (9) "?n <= ?n"
neuper@48763
   585
neuper@42451
   586
ML {*
neuper@37949
   587
val Atools_erls = 
neuper@37949
   588
    append_rls "Atools_erls" e_rls
neuper@41922
   589
               [Calc ("HOL.eq",eval_equal "#equal_"),
neuper@37965
   590
                Thm ("not_true",num_str @{thm not_true}),
neuper@37949
   591
		(*"(~ True) = False"*)
neuper@37965
   592
		Thm ("not_false",num_str @{thm not_false}),
neuper@37949
   593
		(*"(~ False) = True"*)
neuper@37965
   594
		Thm ("and_true",num_str @{thm and_true}),
neuper@37949
   595
		(*"(?a & True) = ?a"*)
neuper@37965
   596
		Thm ("and_false",num_str @{thm and_false}),
neuper@37949
   597
		(*"(?a & False) = False"*)
neuper@37965
   598
		Thm ("or_true",num_str @{thm or_true}),
neuper@37949
   599
		(*"(?a | True) = True"*)
neuper@37965
   600
		Thm ("or_false",num_str @{thm or_false}),
neuper@37949
   601
		(*"(?a | False) = ?a"*)
neuper@37949
   602
               
neuper@37965
   603
		Thm ("rat_leq1",num_str @{thm rat_leq1}),
neuper@37965
   604
		Thm ("rat_leq2",num_str @{thm rat_leq2}),
neuper@37965
   605
		Thm ("rat_leq3",num_str @{thm rat_leq3}),
neuper@37965
   606
                Thm ("refl",num_str @{thm refl}),
neuper@48763
   607
		Thm ("order_refl",num_str @{thm order_refl}),
neuper@37965
   608
		Thm ("radd_left_cancel_le",num_str @{thm radd_left_cancel_le}),
neuper@37949
   609
		
neuper@38045
   610
		Calc ("Orderings.ord_class.less",eval_equ "#less_"),
neuper@38045
   611
		Calc ("Orderings.ord_class.less_eq",eval_equ "#less_equal_"),
neuper@37949
   612
		
neuper@37949
   613
		Calc ("Atools.ident",eval_ident "#ident_"),    
neuper@37949
   614
		Calc ("Atools.is'_const",eval_const "#is_const_"),
neuper@37949
   615
		Calc ("Atools.occurs'_in",eval_occurs_in ""),    
neuper@37949
   616
		Calc ("Tools.matches",eval_matches "")
neuper@37949
   617
		];
neuper@37949
   618
neuper@42451
   619
*}
neuper@42451
   620
neuper@42451
   621
ML {*
neuper@37949
   622
val Atools_crls = 
neuper@37949
   623
    append_rls "Atools_crls" e_rls
neuper@41922
   624
               [Calc ("HOL.eq",eval_equal "#equal_"),
neuper@37965
   625
                Thm ("not_true",num_str @{thm not_true}),
neuper@37965
   626
		Thm ("not_false",num_str @{thm not_false}),
neuper@37965
   627
		Thm ("and_true",num_str @{thm and_true}),
neuper@37965
   628
		Thm ("and_false",num_str @{thm and_false}),
neuper@37965
   629
		Thm ("or_true",num_str @{thm or_true}),
neuper@37965
   630
		Thm ("or_false",num_str @{thm or_false}),
neuper@37949
   631
               
neuper@37965
   632
		Thm ("rat_leq1",num_str @{thm rat_leq1}),
neuper@37965
   633
		Thm ("rat_leq2",num_str @{thm rat_leq2}),
neuper@37965
   634
		Thm ("rat_leq3",num_str @{thm rat_leq3}),
neuper@37965
   635
                Thm ("refl",num_str @{thm refl}),
neuper@48763
   636
		Thm ("order_refl",num_str @{thm order_refl}),
neuper@37965
   637
		Thm ("radd_left_cancel_le",num_str @{thm radd_left_cancel_le}),
neuper@37949
   638
		
neuper@38045
   639
		Calc ("Orderings.ord_class.less",eval_equ "#less_"),
neuper@38045
   640
		Calc ("Orderings.ord_class.less_eq",eval_equ "#less_equal_"),
neuper@37949
   641
		
neuper@37949
   642
		Calc ("Atools.ident",eval_ident "#ident_"),    
neuper@37949
   643
		Calc ("Atools.is'_const",eval_const "#is_const_"),
neuper@37949
   644
		Calc ("Atools.occurs'_in",eval_occurs_in ""),    
neuper@37949
   645
		Calc ("Tools.matches",eval_matches "")
neuper@37949
   646
		];
neuper@37949
   647
neuper@37949
   648
(*val atools_erls = ... waere zu testen ...
neuper@37949
   649
    merge_rls calculate_Atools
neuper@37949
   650
	      (append_rls Atools_erls (*i.A. zu viele rules*)
neuper@37949
   651
			  [Calc ("Atools.ident",eval_ident "#ident_"),    
neuper@37949
   652
			   Calc ("Atools.is'_const",eval_const "#is_const_"),
neuper@37949
   653
			   Calc ("Atools.occurs'_in",
neuper@37949
   654
				 eval_occurs_in "#occurs_in"),    
neuper@37949
   655
			   Calc ("Tools.matches",eval_matches "#matches")
neuper@37949
   656
			   ] (*i.A. zu viele rules*)
neuper@37949
   657
			  );*)
neuper@37949
   658
(* val atools_erls = prep_rls(
neuper@37949
   659
  Rls {id="atools_erls",preconds = [], rew_ord = ("termlessI",termlessI), 
neuper@42451
   660
      erls = e_rls, srls = Erls, calc = [], errpatts = [],
neuper@37965
   661
      rules = [Thm ("refl",num_str @{thm refl}),
neuper@48764
   662
		Thm ("order_refl",num_str @{thm order_refl}),
neuper@37965
   663
		Thm ("radd_left_cancel_le",num_str @{thm radd_left_cancel_le}),
neuper@37965
   664
		Thm ("not_true",num_str @{thm not_true}),
neuper@37965
   665
		Thm ("not_false",num_str @{thm not_false}),
neuper@37965
   666
		Thm ("and_true",num_str @{thm and_true}),
neuper@37965
   667
		Thm ("and_false",num_str @{thm and_false}),
neuper@37965
   668
		Thm ("or_true",num_str @{thm or_true}),
neuper@37965
   669
		Thm ("or_false",num_str @{thm or_false}),
neuper@37965
   670
		Thm ("and_commute",num_str @{thm and_commute}),
neuper@37965
   671
		Thm ("or_commute",num_str @{thm or_commute}),
neuper@37949
   672
		
neuper@38045
   673
		Calc ("Orderings.ord_class.less",eval_equ "#less_"),
neuper@38045
   674
		Calc ("Orderings.ord_class.less_eq",eval_equ "#less_equal_"),
neuper@37949
   675
		
neuper@37949
   676
		Calc ("Atools.ident",eval_ident "#ident_"),    
neuper@37949
   677
		Calc ("Atools.is'_const",eval_const "#is_const_"),
neuper@37949
   678
		Calc ("Atools.occurs'_in",eval_occurs_in ""),    
neuper@37949
   679
		Calc ("Tools.matches",eval_matches "")
neuper@37949
   680
	       ],
neuper@48763
   681
      scr = Prog ((term_of o the o (parse @{theory})) "empty_script")
neuper@37949
   682
      }:rls);
neuper@37968
   683
ruleset' := overwritelth @{theory} 
neuper@37949
   684
		(!ruleset',
neuper@37949
   685
		 [("atools_erls",atools_erls)(*FIXXXME:del with rls.rls'*)
neuper@37949
   686
		  ]);
neuper@37949
   687
*)
neuper@37949
   688
"******* Atools.ML end *******";
neuper@37949
   689
neuper@37949
   690
calclist':= overwritel (!calclist', 
neuper@37949
   691
   [("occurs_in",("Atools.occurs'_in", eval_occurs_in "#occurs_in_")),
neuper@37949
   692
    ("some_occur_in",
neuper@37949
   693
     ("Atools.some'_occur'_in", eval_some_occur_in "#some_occur_in_")),
neuper@37949
   694
    ("is_atom"  ,("Atools.is'_atom",eval_is_atom "#is_atom_")),
neuper@37949
   695
    ("is_even"  ,("Atools.is'_even",eval_is_even "#is_even_")),
neuper@37949
   696
    ("is_const" ,("Atools.is'_const",eval_const "#is_const_")),
neuper@38045
   697
    ("le"       ,("Orderings.ord_class.less"        ,eval_equ "#less_")),
neuper@38045
   698
    ("leq"      ,("Orderings.ord_class.less_eq"       ,eval_equ "#less_equal_")),
neuper@37949
   699
    ("ident"    ,("Atools.ident",eval_ident "#ident_")),
neuper@41922
   700
    ("equal"    ,("HOL.eq",eval_equal "#equal_")),
neuper@38014
   701
    ("PLUS"     ,("Groups.plus_class.plus"        ,eval_binop "#add_")),
neuper@38014
   702
    ("MINUS"    ,("Groups.minus_class.minus",eval_binop "#sub_")),
neuper@38034
   703
    ("TIMES"    ,("Groups.times_class.times"        ,eval_binop "#mult_")),
neuper@38014
   704
    ("DIVIDE"  ,("Rings.inverse_class.divide"  ,eval_cancel "#divide_e")),
neuper@37949
   705
    ("POWER"   ,("Atools.pow"  ,eval_binop "#power_")),
neuper@37949
   706
    ("boollist2sum",("Atools.boollist2sum",eval_boollist2sum ""))
neuper@37949
   707
    ]);
neuper@37949
   708
neuper@37949
   709
val list_rls = prep_rls(
neuper@37949
   710
    merge_rls "list_erls"
neuper@37949
   711
	      (Rls {id="replaced",preconds = [], 
neuper@37949
   712
		    rew_ord = ("termlessI", termlessI),
neuper@37949
   713
		    erls = Rls {id="list_elrs", preconds = [], 
neuper@37949
   714
				rew_ord = ("termlessI",termlessI), 
neuper@37949
   715
				erls = e_rls, 
neuper@42451
   716
				srls = Erls, calc = [], errpatts = [],
neuper@38014
   717
				rules = [Calc ("Groups.plus_class.plus", eval_binop "#add_"),
neuper@38045
   718
					 Calc ("Orderings.ord_class.less",eval_equ "#less_")
neuper@37949
   719
					 (*    ~~~~~~ for nth_Cons_*)
neuper@37949
   720
					 ],
neuper@37949
   721
				scr = EmptyScr},
neuper@42451
   722
		    srls = Erls, calc = [], errpatts = [],
neuper@37949
   723
		    rules = [], scr = EmptyScr})
neuper@37949
   724
	      list_rls);
neuper@37965
   725
ruleset' := overwritelthy @{theory} (!ruleset', [("list_rls", list_rls)]);
neuper@37946
   726
*}
neuper@48764
   727
neuper@37906
   728
end