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