src/Tools/Argo/argo_rewr.ML
author wenzelm
Sat, 06 May 2017 21:28:41 +0200
changeset 65750 7f5556f4b584
parent 64927 a5a09855e424
child 66301 8a6a89d6cf2b
permissions -rw-r--r--
added option verbose for SQL source;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
     1
(*  Title:      Tools/Argo/argo_rewr.ML
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
     2
    Author:     Sascha Boehme
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
     3
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
     4
Bottom-up rewriting of expressions based on rewrite rules and rewrite functions.
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
     5
*)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
     6
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
     7
signature ARGO_REWR =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
     8
sig
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
     9
  (* conversions *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    10
  type conv = Argo_Expr.expr -> Argo_Expr.expr * Argo_Proof.conv
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    11
  val keep: conv
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    12
  val seq: conv list -> conv
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    13
  val args: conv -> conv
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    14
  val rewr: Argo_Proof.rewrite -> Argo_Expr.expr -> conv
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    15
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    16
  (* context *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    17
  type context
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    18
  val context: context
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    19
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    20
  (* rewriting *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    21
  val rewrite: context -> conv
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    22
  val rewrite_top: context -> conv
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    23
  val with_proof: conv -> Argo_Expr.expr * Argo_Proof.proof -> Argo_Proof.context ->
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    24
    (Argo_Expr.expr * Argo_Proof.proof) * Argo_Proof.context
64927
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    25
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    26
  (* normalizations *)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    27
  val nnf: context -> context
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    28
  val norm_prop: context -> context
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    29
  val norm_ite: context -> context
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    30
  val norm_eq: context -> context
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    31
  val norm_add: context -> context
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    32
  val norm_mul: context -> context
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    33
  val norm_arith: context -> context
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    34
end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    35
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    36
structure Argo_Rewr: ARGO_REWR =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    37
struct
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    38
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    39
(* conversions *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    40
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    41
(*
64927
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    42
  Conversions are atomic rewrite steps.
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    43
  For every conversion there is a corresponding inference step.
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    44
*)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    45
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    46
type conv = Argo_Expr.expr -> Argo_Expr.expr * Argo_Proof.conv
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    47
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    48
fun keep e = (e, Argo_Proof.keep_conv)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    49
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    50
fun seq [] e = keep e
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    51
  | seq [cv] e = cv e
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    52
  | seq (cv :: cvs) e =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    53
      let val ((e, c2), c1) = cv e |>> seq cvs
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    54
      in (e, Argo_Proof.mk_then_conv c1 c2) end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    55
64927
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    56
fun on_args f (Argo_Expr.E (k, es)) =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    57
  let val (es, cs) = split_list (f es)
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    58
  in (Argo_Expr.E (k, es), Argo_Proof.mk_args_conv cs) end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    59
64927
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    60
fun args cv e = on_args (map cv) e
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    61
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    62
fun all_args cv k (e as Argo_Expr.E (k', _)) = if k = k' then args (all_args cv k) e else cv e
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    63
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    64
fun rewr r e _ = (e, Argo_Proof.mk_rewr_conv r)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    65
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    66
64927
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    67
(* rewriting result *)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    68
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    69
(*
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    70
  After rewriting an expression, further rewritings might be applicable. The result type is
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    71
  a simple means to control which parts of the rewriting result should be rewritten further.
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    72
  Only the top-most part of a result marked by R constructors is amenable to further rewritings.
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    73
*)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    74
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    75
datatype result =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    76
  E of Argo_Expr.expr |
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    77
  R of Argo_Expr.kind * result list 
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    78
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    79
fun expr_of (E e) = e
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    80
  | expr_of (R (k, ps)) = Argo_Expr.E (k, map expr_of ps)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    81
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    82
fun top_most _ (E _) e = keep e
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    83
  | top_most cv (R (_, ps)) e = seq [on_args (map2 (top_most cv) ps), cv] e
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    84
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    85
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    86
(* context *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    87
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    88
(*
64927
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    89
  The context stores lists of rewritings for every expression kind. A rewriting maps the
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    90
  arguments of an expression with matching kind to an optional rewriting result. Each
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    91
  rewriting might decide whether it is applicable to the given expression arguments and
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    92
  might return no result. The first rewriting that produces a result is applied.
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    93
*)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    94
64927
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    95
structure Kindtab = Table(type key = Argo_Expr.kind val ord = Argo_Expr.kind_ord)
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    96
64927
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    97
type context =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    98
  (int -> Argo_Expr.expr list -> (Argo_Proof.rewrite * result) option) list Kindtab.table
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
    99
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   100
val context: context = Kindtab.empty
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   101
64927
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   102
fun add_func k f = Kindtab.map_default (k, []) (fn fs => fs @ [f])
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   103
fun add_func' k f = add_func k (fn _ => f)
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   104
64927
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   105
fun unary k f = add_func' k (fn [e] => f e | _ => raise Fail "not unary")
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   106
fun binary k f = add_func' k (fn [e1, e2] => f e1 e2 | _ => raise Fail "not binary")
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   107
fun ternary k f = add_func' k (fn [e1, e2, e3] => f e1 e2 e3 | _ => raise Fail "not ternary")
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   108
fun nary k f = add_func k f
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   109
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   110
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   111
(* rewriting *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   112
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   113
(*
64927
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   114
  Rewriting proceeds bottom-up. The top-most result of a rewriting step is rewritten again
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   115
  bottom-up, if necessary. Only the first rewriting that produces a result for a given
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   116
  expression is applied. N-ary expressions are flattened before they are rewritten. For
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   117
  instance, flattening (and p (and q r)) produces (and p q r) where p, q and r are no
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   118
  conjunctions.
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   119
*)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   120
64927
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   121
fun first_rewr cv cx k n es e =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   122
  (case get_first (fn f => f n es) (Kindtab.lookup_list cx k) of
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   123
    NONE => keep e
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   124
  | SOME (r, res) => seq [rewr r (expr_of res), top_most cv res] e) 
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   125
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   126
fun all_args_of k (e as Argo_Expr.E (k', es)) = if k = k' then maps (all_args_of k) es else [e]
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   127
fun kind_depth_of k (Argo_Expr.E (k', es)) =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   128
  if k = k' then 1 + fold (Integer.max o kind_depth_of k) es 0 else 0
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   129
64927
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   130
fun norm_kind cv cx (e as Argo_Expr.E (k, es)) =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   131
  let val (n, es) = if Argo_Expr.is_nary k then (kind_depth_of k e, all_args_of k e) else (1, es)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   132
  in first_rewr cv cx k n es e end
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   133
64927
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   134
fun norm_args cv d (e as Argo_Expr.E (k, _)) =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   135
  if d = 0 then keep e
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   136
  else if Argo_Expr.is_nary k then all_args (cv (d - 1)) k e
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   137
  else args (cv (d - 1)) e
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   138
64927
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   139
fun norm cx d e = seq [norm_args (norm cx) d, norm_kind (norm cx 0) cx] e
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   140
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   141
fun rewrite cx = norm cx ~1   (* bottom-up rewriting *)
64927
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   142
fun rewrite_top cx = norm cx 0   (* top-most rewriting *)
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   143
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   144
fun with_proof cv (e, p) prf =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   145
  let
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   146
    val (e, c) = cv e
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   147
    val (p, prf) = Argo_Proof.mk_rewrite c p prf
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   148
  in ((e, p), prf) end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   149
64927
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   150
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   151
(* result constructors *)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   152
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   153
fun mk_nary _ e [] = e
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   154
  | mk_nary _ _ [e] = e
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   155
  | mk_nary k _ es = R (k, es)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   156
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   157
val e_true = E Argo_Expr.true_expr
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   158
val e_false = E Argo_Expr.false_expr
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   159
fun mk_not e = R (Argo_Expr.Not, [e])
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   160
fun mk_and es = mk_nary Argo_Expr.And e_true es
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   161
fun mk_or es = mk_nary Argo_Expr.Or e_false es
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   162
fun mk_iff e1 e2 = R (Argo_Expr.Iff, [e1, e2])
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   163
fun mk_ite e1 e2 e3 = R (Argo_Expr.Ite, [e1, e2, e3])
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   164
fun mk_num n = E (Argo_Expr.mk_num n)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   165
fun mk_neg e = R (Argo_Expr.Neg, [e])
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   166
fun mk_add [] = raise Fail "bad addition"
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   167
  | mk_add [e] = e
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   168
  | mk_add es = R (Argo_Expr.Add, es)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   169
fun mk_sub e1 e2 = R (Argo_Expr.Sub, [e1, e2])
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   170
fun mk_mul e1 e2 = R (Argo_Expr.Mul, [e1, e2])
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   171
fun mk_div e1 e2 = R (Argo_Expr.Div, [e1, e2])
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   172
fun mk_le e1 e2 = R (Argo_Expr.Le, [e1, e2])
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   173
fun mk_le' e1 e2 = mk_le (E e1) (E e2)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   174
fun mk_eq e1 e2 = R (Argo_Expr.Eq, [e1, e2])
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   175
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   176
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   177
(* rewriting to negation normal form *)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   178
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   179
fun rewr_not (Argo_Expr.E exp) =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   180
  (case exp of
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   181
    (Argo_Expr.True, _) => SOME (Argo_Proof.Rewr_Not_True, e_false)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   182
  | (Argo_Expr.False, _) => SOME (Argo_Proof.Rewr_Not_False, e_true)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   183
  | (Argo_Expr.Not, [e]) => SOME (Argo_Proof.Rewr_Not_Not, E e)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   184
  | (Argo_Expr.And, es) => SOME (Argo_Proof.Rewr_Not_And (length es), mk_or (map (mk_not o E) es))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   185
  | (Argo_Expr.Or, es) => SOME (Argo_Proof.Rewr_Not_Or (length es), mk_and (map (mk_not o E) es))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   186
  | (Argo_Expr.Iff, [Argo_Expr.E (Argo_Expr.Not, [e1]), e2]) =>
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   187
      SOME (Argo_Proof.Rewr_Not_Iff, mk_iff (E e1) (E e2))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   188
  | (Argo_Expr.Iff, [e1, Argo_Expr.E (Argo_Expr.Not, [e2])]) =>
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   189
      SOME (Argo_Proof.Rewr_Not_Iff, mk_iff (E e1) (E e2))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   190
  | (Argo_Expr.Iff, [e1, e2]) => 
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   191
      SOME (Argo_Proof.Rewr_Not_Iff, mk_iff (mk_not (E e1)) (E e2))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   192
  | _ => NONE)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   193
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   194
val nnf = unary Argo_Expr.Not rewr_not
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   195
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   196
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   197
(* propositional normalization *)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   198
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   199
(*
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   200
  Propositional expressions are transformed into literals in the clausifier. Having
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   201
  fewer literals results in faster solver execution. Normalizing propositional expressions
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   202
  turns similar expressions into equal expressions, for which the same literal can be used.
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   203
  The clausifier expects that only negation, disjunction, conjunction and equivalence form
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   204
  propositional expressions. Expressions may be simplified to truth or falsity, but both
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   205
  truth and falsity eventually occur nowhere inside expressions.
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   206
*)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   207
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   208
fun first_index pred xs =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   209
  let val i = find_index pred xs
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   210
  in if i >= 0 then SOME i else NONE end
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   211
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   212
fun rewr_zero r zero _ es =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   213
  Option.map (fn i => (r i, E zero)) (first_index (curry Argo_Expr.eq_expr zero) es)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   214
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   215
fun rewr_dual r zero _ =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   216
  let
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   217
    fun duals _ [] = NONE
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   218
      | duals _ [_] = NONE
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   219
      | duals i (e :: es) =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   220
          (case first_index (Argo_Expr.dual_expr e) es of
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   221
            NONE => duals (i + 1) es
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   222
          | SOME i' => SOME (r (i, i + i' + 1), zero))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   223
  in duals 0 end
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   224
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   225
fun rewr_sort r one mk n es =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   226
  let
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   227
    val l = length es
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   228
    fun add (i, e) = if Argo_Expr.eq_expr (e, one) then I else Argo_Exprtab.cons_list (e, i)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   229
    fun dest (e, i) (es, is) = (e :: es, i :: is)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   230
    val (es, iss) = Argo_Exprtab.fold_rev dest (fold_index add es Argo_Exprtab.empty) ([], [])
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   231
    fun identity is = length is = l andalso forall (op =) (map_index I is)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   232
  in
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   233
    if null iss then SOME (r (l, [[0]]), E one)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   234
    else if n = 1 andalso identity (map hd iss) then NONE
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   235
    else (SOME (r (l, iss), mk (map E es)))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   236
  end
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   237
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   238
fun rewr_imp e1 e2 = SOME (Argo_Proof.Rewr_Imp, mk_or [mk_not (E e1), E e2])
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   239
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   240
fun rewr_iff (e1 as Argo_Expr.E exp1) (e2 as Argo_Expr.E exp2) =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   241
  (case (exp1, exp2) of
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   242
    ((Argo_Expr.True, _), _) => SOME (Argo_Proof.Rewr_Iff_True, E e2)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   243
  | ((Argo_Expr.False, _), _) => SOME (Argo_Proof.Rewr_Iff_False, mk_not (E e2))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   244
  | (_, (Argo_Expr.True, _)) => SOME (Argo_Proof.Rewr_Iff_True, E e1)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   245
  | (_, (Argo_Expr.False, _)) => SOME (Argo_Proof.Rewr_Iff_False, mk_not (E e1))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   246
  | ((Argo_Expr.Not, [e1']), (Argo_Expr.Not, [e2'])) =>
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   247
      SOME (Argo_Proof.Rewr_Iff_Not_Not, mk_iff (E e1') (E e2'))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   248
  | _ =>
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   249
      if Argo_Expr.dual_expr e1 e2 then SOME (Argo_Proof.Rewr_Iff_Dual, e_false)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   250
      else
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   251
        (case Argo_Expr.expr_ord (e1, e2) of
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   252
          EQUAL => SOME (Argo_Proof.Rewr_Iff_Refl, e_true)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   253
        | GREATER => SOME (Argo_Proof.Rewr_Iff_Symm, mk_iff (E e2) (E e1))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   254
        | LESS => NONE))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   255
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   256
val norm_prop =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   257
  nary Argo_Expr.And (rewr_zero Argo_Proof.Rewr_And_False Argo_Expr.false_expr) #>
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   258
  nary Argo_Expr.And (rewr_dual Argo_Proof.Rewr_And_Dual e_false) #>
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   259
  nary Argo_Expr.And (rewr_sort Argo_Proof.Rewr_And_Sort Argo_Expr.true_expr mk_and) #>
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   260
  nary Argo_Expr.Or (rewr_zero Argo_Proof.Rewr_Or_True Argo_Expr.true_expr) #>
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   261
  nary Argo_Expr.Or (rewr_dual Argo_Proof.Rewr_Or_Dual e_true) #>
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   262
  nary Argo_Expr.Or (rewr_sort Argo_Proof.Rewr_Or_Sort Argo_Expr.false_expr mk_or) #>
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   263
  binary Argo_Expr.Imp rewr_imp #>
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   264
  binary Argo_Expr.Iff rewr_iff
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   265
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   266
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   267
(* normalization of if-then-else expressions *)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   268
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   269
fun rewr_ite (Argo_Expr.E (Argo_Expr.True, _)) e _ = SOME (Argo_Proof.Rewr_Ite_True, E e)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   270
  | rewr_ite (Argo_Expr.E (Argo_Expr.False, _)) _ e = SOME (Argo_Proof.Rewr_Ite_False, E e)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   271
  | rewr_ite e1 e2 e3 =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   272
      if Argo_Expr.type_of e2 = Argo_Expr.Bool then
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   273
        SOME (Argo_Proof.Rewr_Ite_Prop,
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   274
          mk_and (map mk_or [[mk_not (E e1), E e2], [E e1, E e3], [E e2, E e3]]))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   275
      else if Argo_Expr.eq_expr (e2, e3) then SOME (Argo_Proof.Rewr_Ite_Eq, E e2)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   276
      else NONE
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   277
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   278
val norm_ite = ternary Argo_Expr.Ite rewr_ite
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   279
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   280
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   281
(* normalization of equality *)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   282
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   283
(*
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   284
  In a normalized equality, the left-hand side is less than the right-hand side with respect to
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   285
  the expression order.
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   286
*)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   287
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   288
fun rewr_eq e1 e2 =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   289
  (case Argo_Expr.expr_ord (e1, e2) of
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   290
    EQUAL => SOME (Argo_Proof.Rewr_Eq_Refl, e_true)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   291
  | GREATER => SOME (Argo_Proof.Rewr_Eq_Symm, mk_eq (E e2) (E e1))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   292
  | LESS => NONE)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   293
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   294
val norm_eq = binary Argo_Expr.Eq rewr_eq
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   295
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   296
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   297
(* arithmetic normalization *)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   298
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   299
(* expression functions *)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   300
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   301
fun scale n e =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   302
  if n = @0 then mk_num @0
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   303
  else if n = @1 then e
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   304
  else mk_mul (mk_num n) e
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   305
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   306
fun dest_factor (Argo_Expr.E (Argo_Expr.Mul, [Argo_Expr.E (Argo_Expr.Num n, _), _])) = n
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   307
  | dest_factor _ = @1
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   308
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   309
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   310
(*
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   311
  Products are normalized either to a number or to the monomial form
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   312
    a * x
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   313
  where a is a non-zero number and is a variable or a product of variables.
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   314
  If x is a product, it contains no number factors. If x is a product, it is sorted
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   315
  based on the expression order. Hence, the product z * y * x will be rewritten
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   316
  to x * y * z. The coefficient a is dropped if it is equal to one;
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   317
  instead of 1 * x the expression x is used.
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   318
*)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   319
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   320
fun mk_mul_comm e1 e2 = (Argo_Proof.Rewr_Mul_Comm, mk_mul (E e2) (E e1))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   321
fun mk_mul_assocr e1 e2 e3 =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   322
  (Argo_Proof.Rewr_Mul_Assoc Argo_Proof.Right, mk_mul (mk_mul (E e1) (E e2)) (E e3))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   323
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   324
  (* commute numbers to the left *)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   325
fun rewr_mul (Argo_Expr.E (Argo_Expr.Num n1, _)) (Argo_Expr.E (Argo_Expr.Num n2, _)) =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   326
      SOME (Argo_Proof.Rewr_Mul_Nums (n1, n2), mk_num (n1 * n2))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   327
  | rewr_mul e1 (e2 as Argo_Expr.E (Argo_Expr.Num _, _)) = SOME (mk_mul_comm e1 e2)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   328
  | rewr_mul e1 (Argo_Expr.E (Argo_Expr.Mul, [e2 as Argo_Expr.E (Argo_Expr.Num _, _), e3])) =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   329
      SOME (mk_mul_assocr e1 e2 e3)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   330
  (* apply distributivity *)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   331
  | rewr_mul (Argo_Expr.E (Argo_Expr.Add, es)) e =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   332
      SOME (Argo_Proof.Rewr_Mul_Sum Argo_Proof.Left, mk_add (map (fn e' => mk_mul (E e') (E e)) es))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   333
  | rewr_mul e (Argo_Expr.E (Argo_Expr.Add, es)) =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   334
      SOME (Argo_Proof.Rewr_Mul_Sum Argo_Proof.Right, mk_add (map (mk_mul (E e) o E) es))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   335
  (* commute non-numerical factors to the right *)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   336
  | rewr_mul (Argo_Expr.E (Argo_Expr.Mul, [e1, e2])) e3 =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   337
      SOME (Argo_Proof.Rewr_Mul_Assoc Argo_Proof.Left, mk_mul (E e1) (mk_mul (E e2) (E e3)))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   338
  (* reduce special products *)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   339
  | rewr_mul (e1 as Argo_Expr.E (Argo_Expr.Num n, _)) e2 =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   340
      if n = @0 then SOME (Argo_Proof.Rewr_Mul_Zero, E e1)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   341
      else if n = @1 then SOME (Argo_Proof.Rewr_Mul_One, E e2)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   342
      else NONE
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   343
  (* combine products with quotients *)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   344
  | rewr_mul (Argo_Expr.E (Argo_Expr.Div, [e1, e2])) e3 =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   345
      SOME (Argo_Proof.Rewr_Mul_Div Argo_Proof.Left, mk_div (mk_mul (E e1) (E e3)) (E e2))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   346
  | rewr_mul e1 (Argo_Expr.E (Argo_Expr.Div, [e2, e3])) =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   347
      SOME (Argo_Proof.Rewr_Mul_Div Argo_Proof.Right, mk_div (mk_mul (E e1) (E e2)) (E e3))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   348
  (* sort non-numerical factors *)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   349
  | rewr_mul e1 (Argo_Expr.E (Argo_Expr.Mul, [e2, e3])) =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   350
      (case Argo_Expr.expr_ord (e1, e2) of
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   351
        GREATER => SOME (mk_mul_assocr e1 e2 e3)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   352
      | _ => NONE)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   353
  | rewr_mul e1 e2 =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   354
      (case Argo_Expr.expr_ord (e1, e2) of
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   355
        GREATER => SOME (mk_mul_comm e1 e2)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   356
      | _ => NONE)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   357
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   358
(*
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   359
  Quotients are normalized either to a number or to the monomial form
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   360
    a * x
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   361
  where a is a non-zero number and x is a variable. If x is a quotient,
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   362
  both dividend and divisor are not a number. The dividend and the divisor may both
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   363
  be products. If so, neither dividend nor divisor contains a numerical factor.
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   364
  Both dividend and divisor are not themselves quotients again. The dividend is never
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   365
  a sum; distributivity is applied to such quotients. The coefficient a is dropped
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   366
  if it is equal to one; instead of 1 * x the expression x is used.
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   367
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   368
  Several non-linear expressions can be rewritten to the described normal form.
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   369
  For example, the expressions (x * z) / y and x * (z / y) will be treated as equal
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   370
  variables by the arithmetic decision procedure. Yet, non-linear expression rewriting
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   371
  is incomplete and keeps several other expressions unchanged.
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   372
*)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   373
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   374
fun rewr_div (Argo_Expr.E (Argo_Expr.Div, [e1, e2])) e3 =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   375
      SOME (Argo_Proof.Rewr_Div_Div Argo_Proof.Left, mk_div (E e1) (mk_mul (E e2) (E e3)))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   376
  | rewr_div e1 (Argo_Expr.E (Argo_Expr.Div, [e2, e3])) =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   377
      SOME (Argo_Proof.Rewr_Div_Div Argo_Proof.Right, mk_div (mk_mul (E e1) (E e3)) (E e2))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   378
  | rewr_div (Argo_Expr.E (Argo_Expr.Num n1, _)) (Argo_Expr.E (Argo_Expr.Num n2, _)) =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   379
      if n2 = @0 then NONE
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   380
      else SOME (Argo_Proof.Rewr_Div_Nums (n1, n2), mk_num (n1 / n2))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   381
  | rewr_div (Argo_Expr.E (Argo_Expr.Num n, _)) e =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   382
      if n = @0 then SOME (Argo_Proof.Rewr_Div_Zero, mk_num @0)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   383
      else if n = @1 then NONE
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   384
      else SOME (Argo_Proof.Rewr_Div_Num (Argo_Proof.Left, n), scale n (mk_div (mk_num @1) (E e)))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   385
  | rewr_div (Argo_Expr.E (Argo_Expr.Mul, [Argo_Expr.E (Argo_Expr.Num n, _), e1])) e2 =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   386
      SOME (Argo_Proof.Rewr_Div_Mul (Argo_Proof.Left, n), scale n (mk_div (E e1) (E e2)))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   387
  | rewr_div e (Argo_Expr.E (Argo_Expr.Num n, _)) =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   388
      if n = @0 then NONE
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   389
      else if n = @1 then SOME (Argo_Proof.Rewr_Div_One, E e)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   390
      else SOME (Argo_Proof.Rewr_Div_Num (Argo_Proof.Right, n), scale (Rat.inv n) (E e))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   391
  | rewr_div e1 (Argo_Expr.E (Argo_Expr.Mul, [Argo_Expr.E (Argo_Expr.Num n, _), e2])) =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   392
      SOME (Argo_Proof.Rewr_Div_Mul (Argo_Proof.Right, n), scale (Rat.inv n) (mk_div (E e1) (E e2)))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   393
  | rewr_div (Argo_Expr.E (Argo_Expr.Add, es)) e =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   394
      SOME (Argo_Proof.Rewr_Div_Sum, mk_add (map (fn e' => mk_div (E e') (E e)) es))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   395
  | rewr_div _ _ = NONE
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   396
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   397
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   398
(*
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   399
  Sums are flattened and normalized to the polynomial form
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   400
    a_0 + a_1 * x_1 + ... + a_n * x_n
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   401
  where all variables x_i are disjoint and where all coefficients a_i are
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   402
  non-zero numbers. Coefficients equal to one are dropped; instead of 1 * x_i
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   403
  the reduced monom x_i is used. The variables x_i are ordered based on the
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   404
  expression order to reduce the number of problem literals by sharing equal
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   405
  expressions.
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   406
*)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   407
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   408
fun add_monom_expr i n e (p, s, etab) =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   409
  let val etab = Argo_Exprtab.map_default (e, (i, @0)) (apsnd (Rat.add n)) etab
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   410
  in ((n, Option.map fst (Argo_Exprtab.lookup etab e)) :: p, s, etab) end
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   411
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   412
fun add_monom (_, Argo_Expr.E (Argo_Expr.Num n, _)) (p, s, etab) = ((n, NONE) :: p, s + n, etab)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   413
  | add_monom (i, Argo_Expr.E (Argo_Expr.Mul, [Argo_Expr.E (Argo_Expr.Num n, _), e])) x =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   414
      add_monom_expr i n e x
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   415
  | add_monom (i, e) x = add_monom_expr i @1 e x
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   416
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   417
fun rewr_add d es =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   418
  let
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   419
    val (p1, s, etab) = fold_index add_monom es ([], @0, Argo_Exprtab.empty)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   420
    val (p2, es) =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   421
      []
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   422
      |> Argo_Exprtab.fold_rev (fn (e, (i, n)) => n <> @0 ? cons ((n, SOME i), scale n (E e))) etab
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   423
      |> s <> @0 ? cons ((s, NONE), mk_num s)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   424
      |> (fn [] => [((@0, NONE), mk_num @0)] | xs => xs)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   425
      |> split_list
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   426
    val ps = (rev p1, p2)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   427
  in
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   428
    if d = 1 andalso eq_list (op =) ps then NONE
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   429
    else SOME (Argo_Proof.Rewr_Add ps, mk_add es)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   430
  end
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   431
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   432
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   433
(*
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   434
  Equations are normalized to the normal form
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   435
    a_0 + a_1 * x_1 + ... + a_n * x_n = b
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   436
  or
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   437
    b = a_0 + a_1 * x_1 + ... + a_n * x_n
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   438
  An equation in normal form is rewritten to a conjunction of two non-strict inequalities. 
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   439
*)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   440
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   441
fun rewr_eq_le e1 e2 = SOME (Argo_Proof.Rewr_Eq_Le, mk_and [mk_le' e1 e2, mk_le' e2 e1])
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   442
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   443
fun rewr_arith_eq (Argo_Expr.E (Argo_Expr.Num n1, _)) (Argo_Expr.E (Argo_Expr.Num n2, _)) =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   444
      let val b = (n1 = n2)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   445
      in SOME (Argo_Proof.Rewr_Eq_Nums b, if b then e_true else e_false) end
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   446
  | rewr_arith_eq (e1 as Argo_Expr.E (Argo_Expr.Num _, _)) e2 = rewr_eq_le e1 e2
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   447
  | rewr_arith_eq e1 (e2 as Argo_Expr.E (Argo_Expr.Num _, _)) = rewr_eq_le e1 e2
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   448
  | rewr_arith_eq e1 e2 = SOME (Argo_Proof.Rewr_Eq_Sub, mk_eq (mk_sub (E e1) (E e2)) (mk_num @0))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   449
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   450
fun is_arith e = member (op =) [Argo_Expr.Real] (Argo_Expr.type_of e)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   451
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   452
fun rewr_eq e1 e2 = if is_arith e1 then rewr_arith_eq e1 e2 else NONE
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   453
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   454
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   455
(*
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   456
  Arithmetic inequalities (less and less-than) are normalized to the normal form
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   457
    a_0 + a_1 * x_1 + ... + a_n * x_n ~ b
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   458
  or
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   459
    b ~ a_0 + a_1 * x_1 + ... + a_n * x_n
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   460
  such that most of the coefficients a_i are positive.
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   461
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   462
  Arithmetic inequalities of the form
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   463
    a * x ~ b
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   464
  or
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   465
    b ~ a * x
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   466
  are normalized to the form
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   467
    x ~ c
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   468
  or
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   469
    c ~ x
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   470
  where c is a number.
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   471
*)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   472
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   473
fun norm_cmp_mul k r f e1 e2 n =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   474
  let val es = if n > @0 then [e1, e2] else [e2, e1]
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   475
  in SOME (Argo_Proof.Rewr_Ineq_Mul (r, n), R (k, f (map (scale n o E) es))) end
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   476
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   477
fun count_factors pred es = fold (fn e => if pred (dest_factor e) then Integer.add 1 else I) es 0
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   478
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   479
fun norm_cmp_swap k r f e1 e2 es =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   480
  let
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   481
    val pos = count_factors (fn n => n > @0) es
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   482
    val neg = count_factors (fn n => n < @0) es
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   483
    val keep = pos > neg orelse (pos = neg andalso dest_factor (hd es) > @0)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   484
  in if keep then NONE else norm_cmp_mul k r f e1 e2 @~1 end
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   485
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   486
fun norm_cmp1 k r f e1 (e2 as Argo_Expr.E (Argo_Expr.Mul, [Argo_Expr.E (Argo_Expr.Num n, _), _])) =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   487
      norm_cmp_mul k r f e1 e2 (Rat.inv n)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   488
  | norm_cmp1 k r f e1 (e2 as Argo_Expr.E (Argo_Expr.Add, Argo_Expr.E (Argo_Expr.Num n, _) :: _)) =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   489
      let fun mk e = mk_add [E e, mk_num (~ n)]
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   490
      in SOME (Argo_Proof.Rewr_Ineq_Add (r, ~ n), R (k, f [mk e1, mk e2])) end
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   491
  | norm_cmp1 k r f e1 (e2 as Argo_Expr.E (Argo_Expr.Add, es)) = norm_cmp_swap k r f e1 e2 es
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   492
  | norm_cmp1 _ _ _ _ _ = NONE
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   493
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   494
fun rewr_cmp _ r pred (Argo_Expr.E (Argo_Expr.Num n1, _)) (Argo_Expr.E (Argo_Expr.Num n2, _)) =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   495
      let val b = pred n1 n2
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   496
      in SOME (Argo_Proof.Rewr_Ineq_Nums (r, b), if b then e_true else e_false) end
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   497
  | rewr_cmp k r _ (e1 as Argo_Expr.E (Argo_Expr.Num _, _)) e2 = norm_cmp1 k r I e1 e2
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   498
  | rewr_cmp k r _ e1 (e2 as Argo_Expr.E (Argo_Expr.Num _, _)) = norm_cmp1 k r rev e2 e1
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   499
  | rewr_cmp k r _ e1 e2 =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   500
      SOME (Argo_Proof.Rewr_Ineq_Sub r, R (k, [mk_sub (E e1) (E e2), mk_num @0]))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   501
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   502
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   503
(*
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   504
  Arithmetic expressions are normalized in order to reduce the number of
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   505
  problem literals. Arithmetically equal expressions are normalized to
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   506
  syntactically equal expressions as much as possible.
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   507
*)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   508
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   509
fun rewr_neg e = SOME (Argo_Proof.Rewr_Neg, scale @~1 (E e))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   510
fun rewr_sub e1 e2 = SOME (Argo_Proof.Rewr_Sub, mk_add [E e1, scale @~1 (E e2)])
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   511
fun rewr_abs e = SOME (Argo_Proof.Rewr_Abs, mk_ite (mk_le (mk_num @0) (E e)) (E e) (mk_neg (E e)))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   512
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   513
fun rewr_min e1 e2 =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   514
  (case Argo_Expr.expr_ord (e1, e2) of
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   515
    LESS => SOME (Argo_Proof.Rewr_Min_Lt, mk_ite (mk_le' e1 e2) (E e1) (E e2))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   516
  | EQUAL => SOME (Argo_Proof.Rewr_Min_Eq, E e1)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   517
  | GREATER => SOME (Argo_Proof.Rewr_Min_Gt, mk_ite (mk_le' e2 e1) (E e2) (E e1)))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   518
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   519
fun rewr_max e1 e2 =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   520
  (case Argo_Expr.expr_ord (e1, e2) of
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   521
    LESS => SOME (Argo_Proof.Rewr_Max_Lt, mk_ite (mk_le' e1 e2) (E e2) (E e1))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   522
  | EQUAL => SOME (Argo_Proof.Rewr_Max_Eq, E e1)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   523
  | GREATER => SOME (Argo_Proof.Rewr_Max_Gt, mk_ite (mk_le' e2 e1) (E e1) (E e2)))
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   524
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   525
val norm_add = nary Argo_Expr.Add rewr_add
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   526
val norm_mul = binary Argo_Expr.Mul rewr_mul
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   527
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   528
val norm_arith =
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   529
  unary Argo_Expr.Neg rewr_neg #>
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   530
  binary Argo_Expr.Sub rewr_sub #>
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   531
  norm_mul #>
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   532
  binary Argo_Expr.Div rewr_div #>
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   533
  norm_add #>
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   534
  binary Argo_Expr.Min rewr_min #>
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   535
  binary Argo_Expr.Max rewr_max #>
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   536
  unary Argo_Expr.Abs rewr_abs #>
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   537
  binary Argo_Expr.Eq rewr_eq #>
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   538
  binary Argo_Expr.Le (rewr_cmp Argo_Expr.Le Argo_Proof.Le Rat.le) #>
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   539
  binary Argo_Expr.Lt (rewr_cmp Argo_Expr.Lt Argo_Proof.Lt Rat.lt)
a5a09855e424 less complex rewriting of Argo expressions: select candidate rewrite rules by analyzing only on the kind of the head expression, then select the applicable rewrite rule using ML pattern matching; keep all normalization code in a single place: modules of decision procedures focus on their core aspects; more normalization for non-linear arithmetic
boehmes
parents: 63960
diff changeset
   540
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   541
end