src/Tools/isac/ProgLang/Atools.thy
author Walther Neuper <wneuper@ist.tugraz.at>
Tue, 03 Apr 2018 14:50:58 +0200
changeset 59424 406681ebe781
parent 59416 src/Tools/isac/Knowledge/Atools.thy@229e5c9cf78b
child 59429 c0fe04973189
permissions -rw-r--r--
partial_function: shift respective thys to ProgLang

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