src/Tools/isac/Knowledge/Atools.thy
author Walther Neuper <neuper@ist.tugraz.at>
Tue, 01 Mar 2011 15:23:59 +0100
branchdecompose-isar
changeset 41905 b772eb34c16c
parent 38053 bb6004e10e71
child 41919 c85b0a1916a5
permissions -rw-r--r--
intermed.update to Isabelle2011: test/../syntax added

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