src/Provers/Arith/fast_lin_arith.ML
author nipkow
Sat, 09 Jan 1999 15:25:44 +0100
changeset 6074 34242451bc91
parent 6062 ede9fea461f5
child 6079 4f7975c74cdf
permissions -rw-r--r--
Added simproc.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
     1
(*  Title:      Provers/Arith/fast_lin_arith.ML
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
     2
    ID:         $Id$
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
     3
    Author:     Tobias Nipkow
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
     4
    Copyright   1998  TU Munich
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
     5
6062
ede9fea461f5 Small mods.
nipkow
parents: 6056
diff changeset
     6
A generic linear arithmetic package.
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
     7
The two tactics provided:
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
     8
    lin_arith_tac:         int -> tactic
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
     9
cut_lin_arith_tac: thms -> int -> tactic
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    10
Only take premises and conclusions into account
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    11
that are already (negated) (in)equations.
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    12
*)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    13
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    14
(*** Data needed for setting up the linear arithmetic package ***)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    15
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    16
signature LIN_ARITH_DATA =
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    17
sig
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    18
  val add_mono_thms:    thm list
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    19
                            (* [| i rel1 j; m rel2 n |] ==> i + m rel3 j + n *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    20
  val conjI:		thm
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    21
  val ccontr:           thm (* (~ P ==> False) ==> P *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    22
  val lessD:            thm (* m < n ==> Suc m <= n *)
6062
ede9fea461f5 Small mods.
nipkow
parents: 6056
diff changeset
    23
  val neqE:             thm (* [| m ~= n; m < n ==> P; n < m ==> P |] ==> P *)
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    24
  val notI:             thm (* (P ==> False) ==> ~ P *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    25
  val not_leD:          thm (* ~(m <= n) ==> Suc n <= m *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    26
  val not_lessD:        thm (* ~(m < n) ==> n < m *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    27
  val sym:		thm (* x = y ==> y = x *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    28
  val decomp: term ->
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    29
             ((term * int)list * int * string * (term * int)list * int)option
6074
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
    30
  val mkEqTrue: thm -> thm
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
    31
  val mk_Trueprop: term -> term
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
    32
  val neg_prop: term -> term * bool
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    33
  val simp: thm -> thm
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
    34
  val is_False: thm -> bool
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
    35
  val is_nat: typ list * term -> bool
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
    36
  val mk_nat_thm: Sign.sg -> term -> thm
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    37
end;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    38
(*
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    39
decomp(`x Rel y') should yield (p,i,Rel,q,j)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    40
   where Rel is one of "<", "~<", "<=", "~<=" and "=" and
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    41
         p/q is the decomposition of the sum terms x/y into a list
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    42
         of summand * multiplicity pairs and a constant summand.
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    43
6074
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
    44
mkEqTrue(P) = `P == True'
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
    45
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
    46
neg_prop(t) = (nt,neg) if t is wrapped up in Trueprop and
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
    47
  nt is the (logically) negated version of t, where the negation
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
    48
  of a negative term is the term itself (no double negation!);
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
    49
  neg = `t is already negated'.
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
    50
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    51
simp must reduce contradictory <= to False.
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    52
   It should also cancel common summands to keep <= reduced;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    53
   otherwise <= can grow to massive proportions.
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
    54
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
    55
is_nat(parameter-types,t) =  t:nat
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
    56
mk_nat_thm(t) = "0 <= t"
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    57
*)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    58
6062
ede9fea461f5 Small mods.
nipkow
parents: 6056
diff changeset
    59
signature FAST_LIN_ARITH =
ede9fea461f5 Small mods.
nipkow
parents: 6056
diff changeset
    60
sig
6074
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
    61
  val lin_arith_prover: Sign.sg -> thm list -> term -> thm option
6062
ede9fea461f5 Small mods.
nipkow
parents: 6056
diff changeset
    62
  val     lin_arith_tac:             int -> tactic
ede9fea461f5 Small mods.
nipkow
parents: 6056
diff changeset
    63
  val cut_lin_arith_tac: thm list -> int -> tactic
ede9fea461f5 Small mods.
nipkow
parents: 6056
diff changeset
    64
end;
ede9fea461f5 Small mods.
nipkow
parents: 6056
diff changeset
    65
ede9fea461f5 Small mods.
nipkow
parents: 6056
diff changeset
    66
functor Fast_Lin_Arith(LA_Data:LIN_ARITH_DATA):FAST_LIN_ARITH =
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    67
struct
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    68
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    69
(*** A fast decision procedure ***)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    70
(*** Code ported from HOL Light ***)
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
    71
(* possible optimizations:
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
    72
   use (var,coeff) rep or vector rep  tp save space;
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
    73
   treat non-negative atoms separately rather than adding 0 <= atom
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
    74
*)
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    75
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    76
datatype lineq_type = Eq | Le | Lt;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    77
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
    78
datatype injust = Asm of int
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
    79
                | Nat of int (* index of atom *)
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    80
                | Fwd of injust * thm
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    81
                | Multiplied of int * injust
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    82
                | Added of injust * injust;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    83
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    84
datatype lineq = Lineq of int * lineq_type * int list * injust;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    85
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    86
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    87
(* Calculate new (in)equality type after addition.                           *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    88
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    89
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    90
fun find_add_type(Eq,x) = x
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    91
  | find_add_type(x,Eq) = x
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    92
  | find_add_type(_,Lt) = Lt
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    93
  | find_add_type(Lt,_) = Lt
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    94
  | find_add_type(Le,Le) = Le;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    95
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    96
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    97
(* Multiply out an (in)equation.                                             *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    98
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    99
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   100
fun multiply_ineq n (i as Lineq(k,ty,l,just)) =
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   101
  if n = 1 then i
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   102
  else if n = 0 andalso ty = Lt then sys_error "multiply_ineq"
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   103
  else if n < 0 andalso (ty=Le orelse ty=Lt) then sys_error "multiply_ineq"
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   104
  else Lineq(n * k,ty,map (apl(n,op * )) l,Multiplied(n,just));
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   105
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   106
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   107
(* Add together (in)equations.                                               *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   108
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   109
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   110
fun add_ineq (i1 as Lineq(k1,ty1,l1,just1)) (i2 as Lineq(k2,ty2,l2,just2)) =
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   111
  let val l = map2 (op +) (l1,l2)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   112
  in Lineq(k1+k2,find_add_type(ty1,ty2),l,Added(just1,just2)) end;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   113
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   114
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   115
(* Elimination of variable between a single pair of (in)equations.           *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   116
(* If they're both inequalities, 1st coefficient must be +ve, 2nd -ve.       *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   117
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   118
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   119
fun gcd x y =
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   120
  let fun gxd x y =
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   121
    if y = 0 then x else gxd y (x mod y)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   122
  in if x < y then gxd y x else gxd x y end;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   123
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   124
fun lcm x y = (x * y) div gcd x y;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   125
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   126
fun el 0 (h::_) = h
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   127
  | el n (_::t) = el (n - 1) t
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   128
  | el _ _  = sys_error "el";
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   129
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   130
fun elim_var v (i1 as Lineq(k1,ty1,l1,just1)) (i2 as Lineq(k2,ty2,l2,just2)) =
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   131
  let val c1 = el v l1 and c2 = el v l2
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   132
      val m = lcm (abs c1) (abs c2)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   133
      val m1 = m div (abs c1) and m2 = m div (abs c2)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   134
      val (n1,n2) =
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   135
        if (c1 >= 0) = (c2 >= 0)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   136
        then if ty1 = Eq then (~m1,m2)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   137
             else if ty2 = Eq then (m1,~m2)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   138
                  else sys_error "elim_var"
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   139
        else (m1,m2)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   140
      val (p1,p2) = if ty1=Eq andalso ty2=Eq andalso (n1 = ~1 orelse n2 = ~1)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   141
                    then (~n1,~n2) else (n1,n2)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   142
  in add_ineq (multiply_ineq n1 i1) (multiply_ineq n2 i2) end;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   143
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   144
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   145
(* The main refutation-finding code.                                         *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   146
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   147
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   148
fun is_trivial (Lineq(_,_,l,_)) = forall (fn i => i=0) l;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   149
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   150
fun is_answer (ans as Lineq(k,ty,l,_)) =
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   151
  case ty  of Eq => k <> 0 | Le => k > 0 | Lt => k >= 0;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   152
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   153
fun calc_blowup l =
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   154
  let val (p,n) = partition (apl(0,op<)) (filter (apl(0,op<>)) l)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   155
  in (length p) * (length n) end;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   156
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   157
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   158
(* Main elimination code:                                                    *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   159
(*                                                                           *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   160
(* (1) Looks for immediate solutions (false assertions with no variables).   *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   161
(*                                                                           *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   162
(* (2) If there are any equations, picks a variable with the lowest absolute *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   163
(* coefficient in any of them, and uses it to eliminate.                     *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   164
(*                                                                           *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   165
(* (3) Otherwise, chooses a variable in the inequality to minimize the       *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   166
(* blowup (number of consequences generated) and eliminates it.              *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   167
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   168
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   169
fun allpairs f xs ys =
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   170
  flat(map (fn x => map (fn y => f x y) ys) xs);
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   171
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   172
fun extract_first p =
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   173
  let fun extract xs (y::ys) = if p y then (Some y,xs@ys)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   174
                               else extract (y::xs) ys
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   175
        | extract xs []      = (None,xs)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   176
  in extract [] end;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   177
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   178
(*
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   179
fun print_ineqs ineqs =
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   180
 writeln(cat_lines(""::map (fn Lineq(c,t,l,_) =>
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   181
   string_of_int c ^
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   182
   (case t of Eq => " =  " | Lt=> " <  " | Le => " <= ") ^
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   183
   commas(map string_of_int l)) ineqs));
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   184
*)
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   185
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   186
fun elim ineqs =
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   187
  let (*val dummy = print_ineqs ineqs;*)
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   188
      val (triv,nontriv) = partition is_trivial ineqs in
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   189
  if not(null triv)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   190
  then case find_first is_answer triv of
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   191
         None => elim nontriv | some => some
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   192
  else
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   193
  if null nontriv then None else
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   194
  let val (eqs,noneqs) = partition (fn (Lineq(_,ty,_,_)) => ty=Eq) nontriv in
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   195
  if not(null eqs) then
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   196
     let val clist = foldl (fn (cs,Lineq(_,_,l,_)) => l union cs) ([],eqs)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   197
         val sclist = sort (fn (x,y) => int_ord(abs(x),abs(y)))
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   198
                           (filter (fn i => i<>0) clist)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   199
         val c = hd sclist
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   200
         val (Some(eq as Lineq(_,_,ceq,_)),othereqs) =
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   201
               extract_first (fn Lineq(_,_,l,_) => c mem l) eqs
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   202
         val v = find_index (fn k => k=c) ceq
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   203
         val (ioth,roth) = partition (fn (Lineq(_,_,l,_)) => el v l = 0)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   204
                                     (othereqs @ noneqs)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   205
         val others = map (elim_var v eq) roth @ ioth
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   206
     in elim others end
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   207
  else
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   208
  let val lists = map (fn (Lineq(_,_,l,_)) => l) noneqs
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   209
      val numlist = 0 upto (length(hd lists) - 1)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   210
      val coeffs = map (fn i => map (el i) lists) numlist
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   211
      val blows = map calc_blowup coeffs
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   212
      val iblows = blows ~~ numlist
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   213
      val nziblows = filter (fn (i,_) => i<>0) iblows
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   214
  in if null nziblows then None else
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   215
     let val (c,v) = hd(sort (fn (x,y) => int_ord(fst(x),fst(y))) nziblows)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   216
         val (no,yes) = partition (fn (Lineq(_,_,l,_)) => el v l = 0) ineqs
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   217
         val (pos,neg) = partition(fn (Lineq(_,_,l,_)) => el v l > 0) yes
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   218
     in elim (no @ allpairs (elim_var v) pos neg) end
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   219
  end
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   220
  end
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   221
  end;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   222
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   223
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   224
(* Translate back a proof.                                                   *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   225
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   226
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   227
(* FIXME OPTIMIZE!!!!
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   228
   Addition/Multiplication need i*t representation rather than t+t+...
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   229
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   230
Simplification may detect a contradiction 'prematurely' due to type
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   231
information: n+1 <= 0 is simplified to False and does not need to be crossed
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   232
with 0 <= n.
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   233
*)
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   234
local
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   235
 exception FalseE of thm
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   236
in
6074
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   237
fun mkthm sg asms just =
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   238
  let val atoms = foldl (fn (ats,(lhs,_,_,rhs,_)) =>
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   239
                            map fst lhs  union  (map fst rhs  union  ats))
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   240
                        ([], mapfilter (LA_Data.decomp o concl_of) asms)
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   241
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   242
      fun addthms thm1 thm2 =
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   243
        let val conj = thm1 RS (thm2 RS LA_Data.conjI)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   244
        in the(get_first (fn th => Some(conj RS th) handle _ => None)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   245
                         LA_Data.add_mono_thms)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   246
        end;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   247
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   248
      fun multn(n,thm) =
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   249
        let fun mul(i,th) = if i=1 then th else mul(i-1, addthms thm th)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   250
        in if n < 0 then mul(~n,thm) RS LA_Data.sym else mul(n,thm) end;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   251
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   252
      fun simp thm =
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   253
        let val thm' = LA_Data.simp thm
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   254
        in if LA_Data.is_False thm' then raise FalseE thm' else thm' end
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   255
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   256
      fun mk(Asm i) = nth_elem(i,asms)
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   257
        | mk(Nat(i)) = LA_Data.mk_nat_thm sg (nth_elem(i,atoms))
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   258
        | mk(Fwd(j,thm)) = mk j RS thm
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   259
        | mk(Added(j1,j2)) = simp(addthms (mk j1) (mk j2))
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   260
        | mk(Multiplied(n,j)) = multn(n,mk j)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   261
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   262
  in LA_Data.simp(mk just) handle FalseE thm => thm end
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   263
end;
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   264
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   265
fun coeff poly atom = case assoc(poly,atom) of None => 0 | Some i => i;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   266
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   267
fun mklineq atoms =
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   268
  let val n = length atoms in
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   269
    fn ((lhs,i,rel,rhs,j),k) =>
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   270
    let val lhsa = map (coeff lhs) atoms
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   271
        and rhsa = map (coeff rhs) atoms
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   272
        val diff = map2 (op -) (rhsa,lhsa)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   273
        val c = i-j
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   274
        val just = Asm k
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   275
    in case rel of
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   276
        "<="   => Some(Lineq(c,Le,diff,just))
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   277
       | "~<=" => Some(Lineq(1-c,Le,map (op ~) diff,Fwd(just,LA_Data.not_leD)))
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   278
       | "<"   => Some(Lineq(c+1,Le,diff,Fwd(just,LA_Data.lessD)))
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   279
       | "~<"  => Some(Lineq(~c,Le,map (op~) diff,Fwd(just,LA_Data.not_lessD)))
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   280
       | "="   => Some(Lineq(c,Eq,diff,just))
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   281
       | "~="  => None
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   282
       | _     => sys_error("mklineq" ^ rel)   
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   283
    end
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   284
  end;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   285
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   286
fun mknat pTs ixs (atom,i) =
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   287
  if LA_Data.is_nat(pTs,atom)
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   288
  then let val l = map (fn j => if j=i then 1 else 0) ixs
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   289
       in Some(Lineq(0,Le,l,Nat(i))) end
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   290
  else None
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   291
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   292
fun abstract pTs items =
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   293
  let val atoms = foldl (fn (ats,((lhs,_,_,rhs,_),_)) =>
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   294
                            (map fst lhs) union ((map fst rhs) union ats))
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   295
                        ([],items)
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   296
      val ixs = 0 upto (length(atoms)-1)
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   297
      val iatoms = atoms ~~ ixs
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   298
  in mapfilter (mklineq atoms) items @ mapfilter (mknat pTs ixs) iatoms end;
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   299
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   300
(* Ordinary refutation *)
6074
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   301
fun refute1(pTs,items) =
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   302
  (case elim (abstract pTs items) of
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   303
       None => []
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   304
     | Some(Lineq(_,_,_,j)) => [j]);
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   305
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   306
fun refute1_tac(i,just) =
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   307
  fn state =>
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   308
    let val sg = #sign(rep_thm state)
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   309
    in resolve_tac [LA_Data.notI,LA_Data.ccontr] i THEN
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   310
       METAHYPS (fn asms => rtac (mkthm sg asms just) 1) i
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   311
    end
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   312
    state;
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   313
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   314
(* Double refutation caused by equality in conclusion *)
6074
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   315
fun refute2(pTs,items, (rhs,i,_,lhs,j), nHs) =
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   316
  (case elim (abstract pTs (items@[((rhs,i,"<",lhs,j),nHs)])) of
6074
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   317
    None => []
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   318
  | Some(Lineq(_,_,_,j1)) =>
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   319
      (case elim (abstract pTs (items@[((lhs,j,"<",rhs,i),nHs)])) of
6074
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   320
        None => []
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   321
      | Some(Lineq(_,_,_,j2)) => [j1,j2]));
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   322
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   323
fun refute2_tac(i,just1,just2) =
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   324
  fn state => 
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   325
    let val sg = #sign(rep_thm state)
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   326
    in rtac LA_Data.ccontr i THEN rotate_tac ~1 i THEN etac LA_Data.neqE i THEN
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   327
       METAHYPS (fn asms => rtac (mkthm sg asms just1) 1) i THEN
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   328
       METAHYPS (fn asms => rtac (mkthm sg asms just2) 1) i
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   329
    end
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   330
    state;
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   331
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   332
fun prove(pTs,Hs,concl) =
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   333
let val nHs = length Hs
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   334
    val ixHs = Hs ~~ (0 upto (nHs-1))
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   335
    val Hitems = mapfilter (fn (h,i) => case LA_Data.decomp h of
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   336
                                 None => None | Some(it) => Some(it,i)) ixHs
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   337
in case LA_Data.decomp concl of
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   338
     None => if null Hitems then [] else refute1(pTs,Hitems)
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   339
   | Some(citem as (r,i,rel,l,j)) =>
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   340
       if rel = "="
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   341
       then refute2(pTs,Hitems,citem,nHs)
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   342
       else let val neg::rel0 = explode rel
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   343
                val nrel = if neg = "~" then implode rel0 else "~"^rel
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   344
            in refute1(pTs, Hitems@[((r,i,nrel,l,j),nHs)]) end
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   345
end;
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   346
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   347
(*
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   348
Fast but very incomplete decider. Only premises and conclusions
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   349
that are already (negated) (in)equations are taken into account.
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   350
*)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   351
val lin_arith_tac = SUBGOAL (fn (A,n) =>
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   352
  let val pTs = rev(map snd (Logic.strip_params A))
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   353
      val Hs = Logic.strip_assums_hyp A
6074
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   354
      val concl = Logic.strip_assums_concl A
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   355
  in case prove(pTs,Hs,concl) of
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   356
       [j] => refute1_tac(n,j)
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   357
     | [j1,j2] => refute2_tac(n,j1,j2)
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   358
     | _ => no_tac
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   359
  end);
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   360
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   361
fun cut_lin_arith_tac thms i = cut_facts_tac thms i THEN lin_arith_tac i;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   362
6074
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   363
fun prover1(just,sg,thms,concl) =
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   364
let val (nconcl,neg) = LA_Data.neg_prop concl
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   365
    val cnconcl = cterm_of sg nconcl
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   366
    val Fthm = mkthm sg (thms @ [assume cnconcl]) just
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   367
    val contr = if neg then LA_Data.notI else LA_Data.ccontr
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   368
in Some(LA_Data.mkEqTrue ((implies_intr cnconcl Fthm) COMP contr)) end
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   369
handle _ => None;
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   370
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   371
(* handle thm with equality conclusion *)
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   372
fun prover2(just1,just2,sg,thms,concl) =
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   373
let val (nconcl,_) = LA_Data.neg_prop concl (* m ~= n *)
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   374
    val cnconcl = cterm_of sg nconcl
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   375
    val neqthm = assume cnconcl
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   376
    val casethm = neqthm COMP LA_Data.neqE (* [|m<n ==> R; n<m ==> R|] ==> R *)
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   377
    val [lessimp1,lessimp2] = prems_of casethm
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   378
    val less1 = fst(Logic.dest_implies lessimp1) (* m<n *)
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   379
    and less2 = fst(Logic.dest_implies lessimp2) (* n<m *)
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   380
    val cless1 = cterm_of sg less1 and cless2 = cterm_of sg less2
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   381
    val thm1 = mkthm sg (thms @ [assume cless1]) just1
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   382
    and thm2 = mkthm sg (thms @ [assume cless2]) just2
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   383
    val dthm1 = implies_intr cless1 thm1 and dthm2 = implies_intr cless2 thm2
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   384
    val thm = dthm2 COMP (dthm1 COMP casethm)
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   385
in Some(LA_Data.mkEqTrue ((implies_intr cnconcl thm) COMP LA_Data.ccontr)) end
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   386
handle _ => None;
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   387
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   388
(* FIXME: forward proof of equalities missing! *)
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   389
fun lin_arith_prover sg thms concl =
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   390
let val Hs = map (#prop o rep_thm) thms
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   391
    val Tconcl = LA_Data.mk_Trueprop concl
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   392
in case prove([],Hs,Tconcl) of
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   393
     [j] => prover1(j,sg,thms,Tconcl)
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   394
   | [j1,j2] => prover2(j1,j2,sg,thms,Tconcl)
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   395
   | _ => None
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   396
end;
6074
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   397
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   398
end;