src/Tools/isac/Knowledge/Atools.thy
author Walther Neuper <neuper@ist.tugraz.at>
Thu, 10 Mar 2011 16:04:00 +0100
branchdecompose-isar
changeset 41928 20138d6136cd
parent 41922 32d7766945fb
child 41972 22c5483e9bfb
permissions -rw-r--r--
intermed.update Isabelle2011: HOL.True

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