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