src/Tools/isac/Knowledge/Atools.thy
author Walther Neuper <neuper@ist.tugraz.at>
Thu, 23 Sep 2010 14:49:23 +0200
branchisac-update-Isa09-2
changeset 38014 3e11e3c2dc42
parent 38011 3147f2c1525c
child 38015 67ba02dffacc
permissions -rw-r--r--
updated "op +", "op -", "op *". "HOL.divide" in src & test

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