src/Provers/Arith/fast_lin_arith.ML
author webertj
Wed, 26 Jul 2006 19:23:04 +0200
changeset 20217 25b068a99d2b
parent 19618 9050a3b01e62
child 20254 58b71535ed00
permissions -rw-r--r--
linear arithmetic splits certain operators (e.g. min, max, abs)
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.
6102
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
     7
It provides two tactics
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
     8
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
     9
    lin_arith_tac:         int -> tactic
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    10
cut_lin_arith_tac: thms -> int -> tactic
6102
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    11
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    12
and a simplification procedure
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    13
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16358
diff changeset
    14
    lin_arith_prover: theory -> simpset -> term -> thm option
6102
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    15
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    16
Only take premises and conclusions into account that are already (negated)
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    17
(in)equations. lin_arith_prover tries to prove or disprove the term.
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    18
*)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    19
9073
40d8dfac96b8 tracing flag for arith_tac
paulson
parents: 8263
diff changeset
    20
(* Debugging: set Fast_Arith.trace *)
7582
2650c9c2ab7f Restructured lin.arith.package.
nipkow
parents: 7570
diff changeset
    21
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    22
(*** 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
    23
6102
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    24
signature LIN_ARITH_LOGIC =
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    25
sig
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
    26
  val conjI:            thm (* P ==> Q ==> P & Q *)
6102
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    27
  val ccontr:           thm (* (~ P ==> False) ==> P *)
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    28
  val notI:             thm (* (P ==> False) ==> ~ P *)
6110
15c2b571225b Simplified interface.
nipkow
parents: 6102
diff changeset
    29
  val not_lessD:        thm (* ~(m < n) ==> n <= m *)
6128
2acc5d36610c More arith refinements.
nipkow
parents: 6110
diff changeset
    30
  val not_leD:          thm (* ~(m <= n) ==> n < m *)
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
    31
  val sym:              thm (* x = y ==> y = x *)
6102
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    32
  val mk_Eq: thm -> thm
16735
008d089822e3 takes & in premises apart now.
nipkow
parents: 16458
diff changeset
    33
  val atomize: thm -> thm list
6102
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    34
  val mk_Trueprop: term -> term
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    35
  val neg_prop: term -> term
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    36
  val is_False: thm -> bool
6128
2acc5d36610c More arith refinements.
nipkow
parents: 6110
diff changeset
    37
  val is_nat: typ list * term -> bool
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16358
diff changeset
    38
  val mk_nat_thm: theory -> term -> thm
6102
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    39
end;
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    40
(*
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    41
mk_Eq(~in) = `in == False'
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    42
mk_Eq(in) = `in == True'
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    43
where `in' is an (in)equality.
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    44
19318
958d5c8a8306 comment fixed
webertj
parents: 19314
diff changeset
    45
neg_prop(t) = neg  if t is wrapped up in Trueprop and
958d5c8a8306 comment fixed
webertj
parents: 19314
diff changeset
    46
  neg is the (logically) negated version of t, where the negation
6102
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    47
  of a negative term is the term itself (no double negation!);
6128
2acc5d36610c More arith refinements.
nipkow
parents: 6110
diff changeset
    48
2acc5d36610c More arith refinements.
nipkow
parents: 6110
diff changeset
    49
is_nat(parameter-types,t) =  t:nat
2acc5d36610c More arith refinements.
nipkow
parents: 6110
diff changeset
    50
mk_nat_thm(t) = "0 <= t"
6102
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    51
*)
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    52
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    53
signature LIN_ARITH_DATA =
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    54
sig
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
    55
  type decompT = (term * Rat.rat) list * Rat.rat * string * (term * Rat.rat) list * Rat.rat * bool  (* internal representation of linear (in-)equations *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
    56
  val decomp: theory -> term -> decompT option
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
    57
  val pre_decomp: theory -> typ list * term list -> (typ list * term list) list  (* preprocessing, performed on a representation of subgoals as list of premises *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
    58
  val pre_tac   : int -> Tactical.tactic                                         (* preprocessing, performed on the goal -- must do the same as 'pre_decomp' *)
16358
2e2a506553a3 IntInf change
nipkow
parents: 15965
diff changeset
    59
  val number_of: IntInf.int * typ -> term
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    60
end;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    61
(*
7551
8e934d1a9ac6 Now distinguishes discrete from non-distrete types.
nipkow
parents: 6128
diff changeset
    62
decomp(`x Rel y') should yield (p,i,Rel,q,j,d)
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    63
   where Rel is one of "<", "~<", "<=", "~<=" and "=" and
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
    64
         p (q, respectively) is the decomposition of the sum term x
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
    65
         (y, respectively) into a list of summand * multiplicity
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
    66
         pairs and a constant summand and d indicates if the domain
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
    67
         is discrete.
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
    68
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
    69
The relationship between pre_decomp and pre_tac is somewhat tricky.  The
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
    70
internal representation of a subgoal and the corresponding theorem must
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
    71
be modified by pre_decomp (pre_tac, resp.) in a corresponding way.  See
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
    72
the comment for split_items below.  (This is even necessary for eta- and
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
    73
beta-equivalent modifications, as some of the lin. arith. code is not
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
    74
insensitive to them.)
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    75
9420
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
    76
ss must reduce contradictory <= to False.
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    77
   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
    78
   otherwise <= can grow to massive proportions.
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    79
*)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    80
6062
ede9fea461f5 Small mods.
nipkow
parents: 6056
diff changeset
    81
signature FAST_LIN_ARITH =
ede9fea461f5 Small mods.
nipkow
parents: 6056
diff changeset
    82
sig
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18572
diff changeset
    83
  val setup: theory -> theory
15184
d2c19aea17bc made mult_mono_thms generic.
nipkow
parents: 15027
diff changeset
    84
  val map_data: ({add_mono_thms: thm list, mult_mono_thms: thm list, inj_thms: thm list,
15922
7ef183f3fc98 neqE applies even if the type is not one which partakes in linear arithmetic.
nipkow
parents: 15660
diff changeset
    85
                 lessD: thm list, neqE: thm list, simpset: Simplifier.simpset}
15184
d2c19aea17bc made mult_mono_thms generic.
nipkow
parents: 15027
diff changeset
    86
                 -> {add_mono_thms: thm list, mult_mono_thms: thm list, inj_thms: thm list,
15922
7ef183f3fc98 neqE applies even if the type is not one which partakes in linear arithmetic.
nipkow
parents: 15660
diff changeset
    87
                     lessD: thm list, neqE: thm list, simpset: Simplifier.simpset})
10575
c78d26d5c3c1 Now adjusted to mixed terms involving coercions.
nipkow
parents: 9420
diff changeset
    88
                -> theory -> theory
19314
cf1c19eee826 comment for conjI added
webertj
parents: 19049
diff changeset
    89
  val trace: bool ref
14510
73ea1234bb23 introduced fast_arith_neq_limit
nipkow
parents: 14386
diff changeset
    90
  val fast_arith_neq_limit: int ref
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16358
diff changeset
    91
  val lin_arith_prover: theory -> simpset -> term -> thm option
17892
62c397c17d18 Simplifier.theory_context;
wenzelm
parents: 17877
diff changeset
    92
  val     lin_arith_tac:    bool -> int -> tactic
17613
072c21e31b42 Simplifier.inherit_bounds;
wenzelm
parents: 17524
diff changeset
    93
  val cut_lin_arith_tac: simpset -> int -> tactic
6062
ede9fea461f5 Small mods.
nipkow
parents: 6056
diff changeset
    94
end;
ede9fea461f5 Small mods.
nipkow
parents: 6056
diff changeset
    95
6102
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    96
functor Fast_Lin_Arith(structure LA_Logic:LIN_ARITH_LOGIC 
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    97
                       and       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
    98
struct
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    99
9420
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   100
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   101
(** theory data **)
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   102
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   103
(* data kind 'Provers/fast_lin_arith' *)
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   104
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16358
diff changeset
   105
structure Data = TheoryDataFun
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16358
diff changeset
   106
(struct
9420
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   107
  val name = "Provers/fast_lin_arith";
15184
d2c19aea17bc made mult_mono_thms generic.
nipkow
parents: 15027
diff changeset
   108
  type T = {add_mono_thms: thm list, mult_mono_thms: thm list, inj_thms: thm list,
15922
7ef183f3fc98 neqE applies even if the type is not one which partakes in linear arithmetic.
nipkow
parents: 15660
diff changeset
   109
            lessD: thm list, neqE: thm list, simpset: Simplifier.simpset};
9420
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   110
10691
4ea37fba9c02 towards rtional arithmetic
nipkow
parents: 10575
diff changeset
   111
  val empty = {add_mono_thms = [], mult_mono_thms = [], inj_thms = [],
15922
7ef183f3fc98 neqE applies even if the type is not one which partakes in linear arithmetic.
nipkow
parents: 15660
diff changeset
   112
               lessD = [], neqE = [], simpset = Simplifier.empty_ss};
9420
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   113
  val copy = I;
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16358
diff changeset
   114
  val extend = I;
9420
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   115
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16358
diff changeset
   116
  fun merge _
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16358
diff changeset
   117
    ({add_mono_thms= add_mono_thms1, mult_mono_thms= mult_mono_thms1, inj_thms= inj_thms1,
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16358
diff changeset
   118
      lessD = lessD1, neqE=neqE1, simpset = simpset1},
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16358
diff changeset
   119
     {add_mono_thms= add_mono_thms2, mult_mono_thms= mult_mono_thms2, inj_thms= inj_thms2,
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16358
diff changeset
   120
      lessD = lessD2, neqE=neqE2, simpset = simpset2}) =
9420
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   121
    {add_mono_thms = Drule.merge_rules (add_mono_thms1, add_mono_thms2),
15184
d2c19aea17bc made mult_mono_thms generic.
nipkow
parents: 15027
diff changeset
   122
     mult_mono_thms = Drule.merge_rules (mult_mono_thms1, mult_mono_thms2),
10575
c78d26d5c3c1 Now adjusted to mixed terms involving coercions.
nipkow
parents: 9420
diff changeset
   123
     inj_thms = Drule.merge_rules (inj_thms1, inj_thms2),
c78d26d5c3c1 Now adjusted to mixed terms involving coercions.
nipkow
parents: 9420
diff changeset
   124
     lessD = Drule.merge_rules (lessD1, lessD2),
15922
7ef183f3fc98 neqE applies even if the type is not one which partakes in linear arithmetic.
nipkow
parents: 15660
diff changeset
   125
     neqE = Drule.merge_rules (neqE1, neqE2),
10575
c78d26d5c3c1 Now adjusted to mixed terms involving coercions.
nipkow
parents: 9420
diff changeset
   126
     simpset = Simplifier.merge_ss (simpset1, simpset2)};
9420
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   127
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   128
  fun print _ _ = ();
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16358
diff changeset
   129
end);
9420
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   130
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   131
val map_data = Data.map;
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18572
diff changeset
   132
val setup = Data.init;
9420
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   133
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   134
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   135
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   136
(*** A fast decision procedure ***)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   137
(*** Code ported from HOL Light ***)
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   138
(* possible optimizations:
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   139
   use (var,coeff) rep or vector rep  tp save space;
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   140
   treat non-negative atoms separately rather than adding 0 <= atom
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   141
*)
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   142
9073
40d8dfac96b8 tracing flag for arith_tac
paulson
parents: 8263
diff changeset
   143
val trace = ref false;
40d8dfac96b8 tracing flag for arith_tac
paulson
parents: 8263
diff changeset
   144
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   145
datatype lineq_type = Eq | Le | Lt;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   146
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   147
datatype injust = Asm of int
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   148
                | Nat of int (* index of atom *)
6128
2acc5d36610c More arith refinements.
nipkow
parents: 6110
diff changeset
   149
                | LessD of injust
2acc5d36610c More arith refinements.
nipkow
parents: 6110
diff changeset
   150
                | NotLessD of injust
2acc5d36610c More arith refinements.
nipkow
parents: 6110
diff changeset
   151
                | NotLeD of injust
7551
8e934d1a9ac6 Now distinguishes discrete from non-distrete types.
nipkow
parents: 6128
diff changeset
   152
                | NotLeDD of injust
16358
2e2a506553a3 IntInf change
nipkow
parents: 15965
diff changeset
   153
                | Multiplied of IntInf.int * injust
2e2a506553a3 IntInf change
nipkow
parents: 15965
diff changeset
   154
                | Multiplied2 of IntInf.int * injust
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   155
                | Added of injust * injust;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   156
16358
2e2a506553a3 IntInf change
nipkow
parents: 15965
diff changeset
   157
datatype lineq = Lineq of IntInf.int * lineq_type * IntInf.int list * injust;
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   158
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   159
fun el 0 (h::_) = h
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   160
  | el n (_::t) = el (n - 1) t
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   161
  | el _ _  = sys_error "el";
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   162
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   163
(* ------------------------------------------------------------------------- *)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   164
(* Finding a (counter) example from the trace of a failed elimination        *)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   165
(* ------------------------------------------------------------------------- *)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   166
(* Examples are represented as rational numbers,                             *)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   167
(* Dont blame John Harrison for this code - it is entirely mine. TN          *)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   168
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   169
exception NoEx;
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   170
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   171
(* Coding: (i,true,cs) means i <= cs and (i,false,cs) means i < cs.
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   172
   In general, true means the bound is included, false means it is excluded.
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   173
   Need to know if it is a lower or upper bound for unambiguous interpretation!
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   174
*)
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   175
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   176
fun elim_eqns(ineqs,Lineq(i,Le,cs,_)) = (i,true,cs)::ineqs
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   177
  | elim_eqns(ineqs,Lineq(i,Eq,cs,_)) = (i,true,cs)::(~i,true,map ~ cs)::ineqs
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   178
  | elim_eqns(ineqs,Lineq(i,Lt,cs,_)) = (i,false,cs)::ineqs;
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   179
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   180
(* PRE: ex[v] must be 0! *)
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17892
diff changeset
   181
fun eval (ex:Rat.rat list) v (a:IntInf.int,le,cs:IntInf.int list) =
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17892
diff changeset
   182
  let val rs = map Rat.rat_of_intinf cs
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17892
diff changeset
   183
      val rsum = Library.foldl Rat.add (Rat.zero, map Rat.mult (rs ~~ ex))
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17892
diff changeset
   184
  in (Rat.mult (Rat.add(Rat.rat_of_intinf a,Rat.neg rsum), Rat.inv(el v rs)), le) end;
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   185
(* If el v rs < 0, le should be negated.
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   186
   Instead this swap is taken into account in ratrelmin2.
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   187
*)
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   188
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   189
fun ratrelmin2(x as (r,ler),y as (s,les)) =
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17892
diff changeset
   190
  if r=s then (r, (not ler) andalso (not les)) else if Rat.le(r,s) then x else y;
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   191
fun ratrelmax2(x as (r,ler),y as (s,les)) =
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17892
diff changeset
   192
  if r=s then (r,ler andalso les) else if Rat.le(r,s) then y else x;
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   193
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   194
val ratrelmin = foldr1 ratrelmin2;
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   195
val ratrelmax = foldr1 ratrelmax2;
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   196
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   197
fun ratexact up (r,exact) =
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   198
  if exact then r else
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17892
diff changeset
   199
  let val (p,q) = Rat.quotient_of_rat r
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17892
diff changeset
   200
      val nth = Rat.inv(Rat.rat_of_intinf q)
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17892
diff changeset
   201
  in Rat.add(r,if up then nth else Rat.neg nth) end;
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   202
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17892
diff changeset
   203
fun ratmiddle(r,s) = Rat.mult(Rat.add(r,s),Rat.inv(Rat.rat_of_int 2));
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   204
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   205
fun choose2 d ((lb, exactl), (ub, exactu)) =
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   206
  if Rat.le (lb, Rat.zero) andalso (lb <> Rat.zero orelse exactl) andalso
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   207
     Rat.le (Rat.zero, ub) andalso (ub <> Rat.zero orelse exactu)
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17892
diff changeset
   208
  then Rat.zero else
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   209
  if not d
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17892
diff changeset
   210
  then (if Rat.ge0 lb
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   211
        then if exactl then lb else ratmiddle (lb, ub)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   212
        else if exactu then ub else ratmiddle (lb, ub))
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   213
  else (* discrete domain, both bounds must be exact *)
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17892
diff changeset
   214
  if Rat.ge0 lb then let val lb' = Rat.roundup lb
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   215
                    in if Rat.le (lb', ub) then lb' else raise NoEx end
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17892
diff changeset
   216
               else let val ub' = Rat.rounddown ub
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   217
                    in if Rat.le (lb, ub') then ub' else raise NoEx end;
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   218
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   219
fun findex1 discr (ex, (v, lineqs)) =
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   220
  let val nz = List.filter (fn (Lineq(_,_,cs,_)) => el v cs <> 0) lineqs;
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   221
      val ineqs = Library.foldl elim_eqns ([],nz)
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   222
      val (ge,le) = List.partition (fn (_,_,cs) => el v cs > 0) ineqs
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   223
      val lb = ratrelmax (map (eval ex v) ge)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   224
      val ub = ratrelmin (map (eval ex v) le)
18011
685d95c793ff cleaned up nth, nth_update, nth_map and nth_string functions
haftmann
parents: 17951
diff changeset
   225
  in nth_update (v, choose2 (nth discr v) (lb, ub)) ex end;
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   226
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   227
fun findex discr = Library.foldl (findex1 discr);
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   228
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   229
fun elim1 v x =
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   230
  map (fn (a,le,bs) => (Rat.add (a, Rat.neg (Rat.mult (el v bs, x))), le,
18011
685d95c793ff cleaned up nth, nth_update, nth_map and nth_string functions
haftmann
parents: 17951
diff changeset
   231
                        nth_update (v, Rat.zero) bs));
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   232
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17892
diff changeset
   233
fun single_var v (_,_,cs) = (filter_out (equal Rat.zero) cs = [el v cs]);
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   234
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   235
(* The base case:
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   236
   all variables occur only with positive or only with negative coefficients *)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   237
fun pick_vars discr (ineqs,ex) =
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17892
diff changeset
   238
  let val nz = filter_out (fn (_,_,cs) => forall (equal Rat.zero) cs) ineqs
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   239
  in case nz of [] => ex
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   240
     | (_,_,cs) :: _ =>
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17892
diff changeset
   241
       let val v = find_index (not o equal Rat.zero) cs
18011
685d95c793ff cleaned up nth, nth_update, nth_map and nth_string functions
haftmann
parents: 17951
diff changeset
   242
           val d = nth discr v
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17892
diff changeset
   243
           val pos = Rat.ge0(el v cs)
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   244
           val sv = List.filter (single_var v) nz
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   245
           val minmax =
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17892
diff changeset
   246
             if pos then if d then Rat.roundup o fst o ratrelmax
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   247
                         else ratexact true o ratrelmax
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17892
diff changeset
   248
                    else if d then Rat.rounddown o fst o ratrelmin
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   249
                         else ratexact false o ratrelmin
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17892
diff changeset
   250
           val bnds = map (fn (a,le,bs) => (Rat.mult(a,Rat.inv(el v bs)),le)) sv
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17892
diff changeset
   251
           val x = minmax((Rat.zero,if pos then true else false)::bnds)
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   252
           val ineqs' = elim1 v x nz
18011
685d95c793ff cleaned up nth, nth_update, nth_map and nth_string functions
haftmann
parents: 17951
diff changeset
   253
           val ex' = nth_update (v, x) ex
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   254
       in pick_vars discr (ineqs',ex') end
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   255
  end;
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   256
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   257
fun findex0 discr n lineqs =
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   258
  let val ineqs = Library.foldl elim_eqns ([],lineqs)
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17892
diff changeset
   259
      val rineqs = map (fn (a,le,cs) => (Rat.rat_of_intinf a, le, map Rat.rat_of_intinf cs))
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   260
                       ineqs
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17892
diff changeset
   261
  in pick_vars discr (rineqs,replicate n Rat.zero) end;
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   262
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   263
(* ------------------------------------------------------------------------- *)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   264
(* End of counter example finder. The actual decision procedure starts here. *)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   265
(* ------------------------------------------------------------------------- *)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   266
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   267
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   268
(* 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
   269
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   270
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   271
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
   272
  | find_add_type(x,Eq) = x
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   273
  | find_add_type(_,Lt) = Lt
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   274
  | find_add_type(Lt,_) = Lt
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   275
  | find_add_type(Le,Le) = Le;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   276
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   277
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   278
(* Multiply out an (in)equation.                                             *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   279
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   280
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   281
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
   282
  if n = 1 then i
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   283
  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
   284
  else if n < 0 andalso (ty=Le orelse ty=Lt) then sys_error "multiply_ineq"
17524
42d56a6dec6e fixed syntax for sml/nj
paulson
parents: 17515
diff changeset
   285
  else Lineq (n * k, ty, map (curry op* n) l, Multiplied (n, just));
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   286
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   287
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   288
(* Add together (in)equations.                                               *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   289
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   290
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   291
fun add_ineq (i1 as Lineq(k1,ty1,l1,just1)) (i2 as Lineq(k2,ty2,l2,just2)) =
18330
444f16d232a2 introduced new map2, fold
haftmann
parents: 18011
diff changeset
   292
  let val l = map2 (curry (op +)) l1 l2
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   293
  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
   294
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   295
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   296
(* 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
   297
(* 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
   298
(* ------------------------------------------------------------------------- *)
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
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
   301
  let val c1 = el v l1 and c2 = el v l2
16358
2e2a506553a3 IntInf change
nipkow
parents: 15965
diff changeset
   302
      val m = lcm(abs c1, abs c2)
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   303
      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
   304
      val (n1,n2) =
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   305
        if (c1 >= 0) = (c2 >= 0)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   306
        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
   307
             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
   308
                  else sys_error "elim_var"
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   309
        else (m1,m2)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   310
      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
   311
                    then (~n1,~n2) else (n1,n2)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   312
  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
   313
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   314
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   315
(* The main refutation-finding code.                                         *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   316
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   317
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   318
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
   319
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   320
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
   321
  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
   322
16358
2e2a506553a3 IntInf change
nipkow
parents: 15965
diff changeset
   323
fun calc_blowup (l:IntInf.int list) =
17496
26535df536ae slight adaptions to library changes
haftmann
parents: 17325
diff changeset
   324
  let val (p,n) = List.partition (curry (op <) 0) (List.filter (curry (op <>) 0) l)
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   325
  in (length p) * (length n) end;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   326
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   327
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   328
(* Main elimination code:                                                    *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   329
(*                                                                           *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   330
(* (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
   331
(*                                                                           *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   332
(* (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
   333
(* 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
   334
(*                                                                           *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   335
(* (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
   336
(* 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
   337
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   338
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   339
fun allpairs f xs ys =
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   340
  List.concat (map (fn x => map (fn y => f x y) ys) xs);
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   341
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   342
fun extract_first p =
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15184
diff changeset
   343
  let fun extract xs (y::ys) = if p y then (SOME y,xs@ys)
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   344
                               else extract (y::xs) ys
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15184
diff changeset
   345
        | extract xs []      = (NONE,xs)
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   346
  in extract [] end;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   347
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   348
fun print_ineqs ineqs =
9073
40d8dfac96b8 tracing flag for arith_tac
paulson
parents: 8263
diff changeset
   349
  if !trace then
12262
11ff5f47df6e use tracing function for trace output;
wenzelm
parents: 12109
diff changeset
   350
     tracing(cat_lines(""::map (fn Lineq(c,t,l,_) =>
16358
2e2a506553a3 IntInf change
nipkow
parents: 15965
diff changeset
   351
       IntInf.toString c ^
9073
40d8dfac96b8 tracing flag for arith_tac
paulson
parents: 8263
diff changeset
   352
       (case t of Eq => " =  " | Lt=> " <  " | Le => " <= ") ^
16358
2e2a506553a3 IntInf change
nipkow
parents: 15965
diff changeset
   353
       commas(map IntInf.toString l)) ineqs))
9073
40d8dfac96b8 tracing flag for arith_tac
paulson
parents: 8263
diff changeset
   354
  else ();
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   355
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   356
type history = (int * lineq list) list;
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   357
datatype result = Success of injust | Failure of history;
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   358
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   359
fun elim (ineqs, hist) =
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   360
  let val dummy = print_ineqs ineqs
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   361
      val (triv, nontriv) = List.partition is_trivial ineqs in
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   362
  if not (null triv)
13186
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   363
  then case Library.find_first is_answer triv of
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   364
         NONE => elim (nontriv, hist)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15184
diff changeset
   365
       | SOME(Lineq(_,_,_,j)) => Success j
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   366
  else
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   367
  if null nontriv then Failure hist
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   368
  else
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   369
  let val (eqs, noneqs) = List.partition (fn (Lineq(_,ty,_,_)) => ty=Eq) nontriv in
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   370
  if not (null eqs) then
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   371
     let val clist = Library.foldl (fn (cs,Lineq(_,_,l,_)) => l union cs) ([],eqs)
16358
2e2a506553a3 IntInf change
nipkow
parents: 15965
diff changeset
   372
         val sclist = sort (fn (x,y) => IntInf.compare(abs(x),abs(y)))
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   373
                           (List.filter (fn i => i<>0) clist)
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   374
         val c = hd sclist
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15184
diff changeset
   375
         val (SOME(eq as Lineq(_,_,ceq,_)),othereqs) =
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   376
               extract_first (fn Lineq(_,_,l,_) => c mem l) eqs
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   377
         val v = find_index_eq c ceq
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   378
         val (ioth,roth) = List.partition (fn (Lineq(_,_,l,_)) => el v l = 0)
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   379
                                     (othereqs @ noneqs)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   380
         val others = map (elim_var v eq) roth @ ioth
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   381
     in elim(others,(v,nontriv)::hist) end
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   382
  else
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   383
  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
   384
      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
   385
      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
   386
      val blows = map calc_blowup coeffs
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   387
      val iblows = blows ~~ numlist
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   388
      val nziblows = List.filter (fn (i,_) => i<>0) iblows
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   389
  in if null nziblows then Failure((~1,nontriv)::hist)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   390
     else
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   391
     let val (c,v) = hd(sort (fn (x,y) => int_ord(fst(x),fst(y))) nziblows)
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   392
         val (no,yes) = List.partition (fn (Lineq(_,_,l,_)) => el v l = 0) ineqs
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   393
         val (pos,neg) = List.partition(fn (Lineq(_,_,l,_)) => el v l > 0) yes
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   394
     in elim(no @ allpairs (elim_var v) pos neg, (v,nontriv)::hist) end
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   395
  end
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   396
  end
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   397
  end;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   398
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   399
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   400
(* Translate back a proof.                                                   *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   401
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   402
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   403
(* string -> Thm.thm -> Thm.thm *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   404
fun trace_thm msg th =
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   405
    (if !trace then (tracing msg; tracing (Display.string_of_thm th)) else (); th);
9073
40d8dfac96b8 tracing flag for arith_tac
paulson
parents: 8263
diff changeset
   406
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   407
(* string -> unit *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   408
fun trace_msg msg =
12262
11ff5f47df6e use tracing function for trace output;
wenzelm
parents: 12109
diff changeset
   409
    if !trace then tracing msg else ();
9073
40d8dfac96b8 tracing flag for arith_tac
paulson
parents: 8263
diff changeset
   410
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   411
(* FIXME OPTIMIZE!!!! (partly done already)
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   412
   Addition/Multiplication need i*t representation rather than t+t+...
10691
4ea37fba9c02 towards rtional arithmetic
nipkow
parents: 10575
diff changeset
   413
   Get rid of Mulitplied(2). For Nat LA_Data.number_of should return Suc^n
4ea37fba9c02 towards rtional arithmetic
nipkow
parents: 10575
diff changeset
   414
   because Numerals are not known early enough.
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   415
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   416
Simplification may detect a contradiction 'prematurely' due to type
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   417
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
   418
with 0 <= n.
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   419
*)
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   420
local
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   421
 exception FalseE of thm
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   422
in
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   423
(* Theory.theory * MetaSimplifier.simpset -> Thm.thm list -> injust -> Thm.thm *)
17515
830bc15e692c Simplifier.inherit_bounds;
wenzelm
parents: 17496
diff changeset
   424
fun mkthm (sg, ss) asms just =
15922
7ef183f3fc98 neqE applies even if the type is not one which partakes in linear arithmetic.
nipkow
parents: 15660
diff changeset
   425
  let val {add_mono_thms, mult_mono_thms, inj_thms, lessD, simpset, ...} =
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16358
diff changeset
   426
          Data.get sg;
17877
67d5ab1cb0d8 Simplifier.inherit_context instead of Simplifier.inherit_bounds;
wenzelm
parents: 17613
diff changeset
   427
      val simpset' = Simplifier.inherit_context ss simpset;
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   428
      val atoms = Library.foldl (fn (ats, (lhs,_,_,rhs,_,_)) =>
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   429
                            map fst lhs  union  (map fst rhs  union  ats))
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   430
                        ([], List.mapPartial (fn thm => if Thm.no_prems thm
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   431
                                              then LA_Data.decomp sg (concl_of thm)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   432
                                              else NONE) asms)
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   433
10575
c78d26d5c3c1 Now adjusted to mixed terms involving coercions.
nipkow
parents: 9420
diff changeset
   434
      fun add2 thm1 thm2 =
6102
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
   435
        let val conj = thm1 RS (thm2 RS LA_Logic.conjI)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15184
diff changeset
   436
        in get_first (fn th => SOME(conj RS th) handle THM _ => NONE) add_mono_thms
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   437
        end;
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15184
diff changeset
   438
      fun try_add [] _ = NONE
10575
c78d26d5c3c1 Now adjusted to mixed terms involving coercions.
nipkow
parents: 9420
diff changeset
   439
        | try_add (thm1::thm1s) thm2 = case add2 thm1 thm2 of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15184
diff changeset
   440
             NONE => try_add thm1s thm2 | some => some;
10575
c78d26d5c3c1 Now adjusted to mixed terms involving coercions.
nipkow
parents: 9420
diff changeset
   441
c78d26d5c3c1 Now adjusted to mixed terms involving coercions.
nipkow
parents: 9420
diff changeset
   442
      fun addthms thm1 thm2 =
c78d26d5c3c1 Now adjusted to mixed terms involving coercions.
nipkow
parents: 9420
diff changeset
   443
        case add2 thm1 thm2 of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15184
diff changeset
   444
          NONE => (case try_add ([thm1] RL inj_thms) thm2 of
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   445
                     NONE => ( the (try_add ([thm2] RL inj_thms) thm1)
15660
255055554c67 handle Option instead of OPTION;
wenzelm
parents: 15570
diff changeset
   446
                               handle Option =>
14360
e654599b114e Added an exception handler and error msg.
nipkow
parents: 13561
diff changeset
   447
                               (trace_thm "" thm1; trace_thm "" thm2;
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   448
                                sys_error "Lin.arith. failed to add thms")
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   449
                             )
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15184
diff changeset
   450
                   | SOME thm => thm)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15184
diff changeset
   451
        | SOME thm => thm;
10575
c78d26d5c3c1 Now adjusted to mixed terms involving coercions.
nipkow
parents: 9420
diff changeset
   452
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   453
      fun multn(n,thm) =
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   454
        let fun mul(i,th) = if i=1 then th else mul(i-1, addthms thm th)
6102
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
   455
        in if n < 0 then mul(~n,thm) RS LA_Logic.sym else mul(n,thm) end;
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   456
15184
d2c19aea17bc made mult_mono_thms generic.
nipkow
parents: 15027
diff changeset
   457
      fun multn2(n,thm) =
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15184
diff changeset
   458
        let val SOME(mth) =
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15184
diff changeset
   459
              get_first (fn th => SOME(thm RS th) handle THM _ => NONE) mult_mono_thms
15184
d2c19aea17bc made mult_mono_thms generic.
nipkow
parents: 15027
diff changeset
   460
            fun cvar(th,_ $ (_ $ _ $ var)) = cterm_of (#sign(rep_thm th)) var;
d2c19aea17bc made mult_mono_thms generic.
nipkow
parents: 15027
diff changeset
   461
            val cv = cvar(mth, hd(prems_of mth));
d2c19aea17bc made mult_mono_thms generic.
nipkow
parents: 15027
diff changeset
   462
            val ct = cterm_of sg (LA_Data.number_of(n,#T(rep_cterm cv)))
d2c19aea17bc made mult_mono_thms generic.
nipkow
parents: 15027
diff changeset
   463
        in instantiate ([],[(cv,ct)]) mth end
10691
4ea37fba9c02 towards rtional arithmetic
nipkow
parents: 10575
diff changeset
   464
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   465
      fun simp thm =
17515
830bc15e692c Simplifier.inherit_bounds;
wenzelm
parents: 17496
diff changeset
   466
        let val thm' = trace_thm "Simplified:" (full_simplify simpset' thm)
6102
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
   467
        in if LA_Logic.is_False thm' then raise FalseE thm' else thm' end
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   468
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   469
      fun mk (Asm i) = trace_thm "Asm" (nth asms i)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   470
        | mk (Nat i) = trace_thm "Nat" (LA_Logic.mk_nat_thm sg (nth atoms i))
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   471
        | mk (LessD(j)) = trace_thm "L" (hd([mk j] RL lessD))
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   472
        | mk (NotLeD(j)) = trace_thm "NLe" (mk j RS LA_Logic.not_leD)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   473
        | mk (NotLeDD(j)) = trace_thm "NLeD" (hd([mk j RS LA_Logic.not_leD] RL lessD))
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   474
        | mk (NotLessD(j)) = trace_thm "NL" (mk j RS LA_Logic.not_lessD)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   475
        | mk (Added(j1,j2)) = simp (trace_thm "+" (addthms (mk j1) (mk j2)))
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   476
        | mk (Multiplied(n,j)) = (trace_msg("*"^IntInf.toString n); trace_thm "*" (multn(n,mk j)))
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   477
        | mk (Multiplied2(n,j)) = simp (trace_msg("**"^IntInf.toString n); trace_thm "**" (multn2(n,mk j)))
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   478
9073
40d8dfac96b8 tracing flag for arith_tac
paulson
parents: 8263
diff changeset
   479
  in trace_msg "mkthm";
12932
3bda5306d262 updated debugging output
nipkow
parents: 12311
diff changeset
   480
     let val thm = trace_thm "Final thm:" (mk just)
17515
830bc15e692c Simplifier.inherit_bounds;
wenzelm
parents: 17496
diff changeset
   481
     in let val fls = simplify simpset' thm
13186
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   482
        in trace_thm "After simplification:" fls;
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   483
           if LA_Logic.is_False fls then fls
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   484
           else
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   485
            (tracing "Assumptions:"; List.app (tracing o Display.string_of_thm) asms;
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   486
             tracing "Proved:"; tracing (Display.string_of_thm fls);
13186
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   487
             warning "Linear arithmetic should have refuted the assumptions.\n\
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   488
                     \Please inform Tobias Nipkow (nipkow@in.tum.de).";
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   489
             fls)
12932
3bda5306d262 updated debugging output
nipkow
parents: 12311
diff changeset
   490
        end
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   491
     end handle FalseE thm => trace_thm "False reached early:" thm
12932
3bda5306d262 updated debugging output
nipkow
parents: 12311
diff changeset
   492
  end
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   493
end;
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   494
16358
2e2a506553a3 IntInf change
nipkow
parents: 15965
diff changeset
   495
fun coeff poly atom : IntInf.int =
17325
d9d50222808e introduced new-style AList operations
haftmann
parents: 17029
diff changeset
   496
  AList.lookup (op =) poly atom |> the_default 0;
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   497
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   498
(* int list -> int *)
16358
2e2a506553a3 IntInf change
nipkow
parents: 15965
diff changeset
   499
fun lcms is = Library.foldl lcm (1, is);
10691
4ea37fba9c02 towards rtional arithmetic
nipkow
parents: 10575
diff changeset
   500
4ea37fba9c02 towards rtional arithmetic
nipkow
parents: 10575
diff changeset
   501
fun integ(rlhs,r,rel,rrhs,s,d) =
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17892
diff changeset
   502
let val (rn,rd) = Rat.quotient_of_rat r and (sn,sd) = Rat.quotient_of_rat s
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17892
diff changeset
   503
    val m = lcms(map (abs o snd o Rat.quotient_of_rat) (r :: s :: map snd rlhs @ map snd rrhs))
15965
f422f8283491 Use of IntInf.int instead of int in most numeric simprocs; avoids
paulson
parents: 15922
diff changeset
   504
    fun mult(t,r) = 
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17892
diff changeset
   505
        let val (i,j) = Rat.quotient_of_rat r
15965
f422f8283491 Use of IntInf.int instead of int in most numeric simprocs; avoids
paulson
parents: 15922
diff changeset
   506
        in (t,i * (m div j)) end
12932
3bda5306d262 updated debugging output
nipkow
parents: 12311
diff changeset
   507
in (m,(map mult rlhs, rn*(m div rd), rel, map mult rrhs, sn*(m div sd), d)) end
10691
4ea37fba9c02 towards rtional arithmetic
nipkow
parents: 10575
diff changeset
   508
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   509
fun mklineq n atoms =
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   510
  fn (item, k) =>
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   511
  let val (m, (lhs,i,rel,rhs,j,discrete)) = integ item
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   512
      val lhsa = map (coeff lhs) atoms
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   513
      and rhsa = map (coeff rhs) atoms
18330
444f16d232a2 introduced new map2, fold
haftmann
parents: 18011
diff changeset
   514
      val diff = map2 (curry (op -)) rhsa lhsa
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   515
      val c = i-j
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   516
      val just = Asm k
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   517
      fun lineq(c,le,cs,j) = Lineq(c,le,cs, if m=1 then j else Multiplied2(m,j))
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   518
  in case rel of
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   519
      "<="   => lineq(c,Le,diff,just)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   520
     | "~<=" => if discrete
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   521
                then lineq(1-c,Le,map (op ~) diff,NotLeDD(just))
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   522
                else lineq(~c,Lt,map (op ~) diff,NotLeD(just))
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   523
     | "<"   => if discrete
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   524
                then lineq(c+1,Le,diff,LessD(just))
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   525
                else lineq(c,Lt,diff,just)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   526
     | "~<"  => lineq(~c,Le,map (op~) diff,NotLessD(just))
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   527
     | "="   => lineq(c,Eq,diff,just)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   528
     | _     => sys_error("mklineq" ^ rel)   
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   529
  end;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   530
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   531
(* ------------------------------------------------------------------------- *)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   532
(* Print (counter) example                                                   *)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   533
(* ------------------------------------------------------------------------- *)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   534
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   535
fun print_atom((a,d),r) =
17951
ff954cc338c7 introduced functions from Pure/General/rat.ML
haftmann
parents: 17892
diff changeset
   536
  let val (p,q) = Rat.quotient_of_rat r
15965
f422f8283491 Use of IntInf.int instead of int in most numeric simprocs; avoids
paulson
parents: 15922
diff changeset
   537
      val s = if d then IntInf.toString p else
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   538
              if p = 0 then "0"
15965
f422f8283491 Use of IntInf.int instead of int in most numeric simprocs; avoids
paulson
parents: 15922
diff changeset
   539
              else IntInf.toString p ^ "/" ^ IntInf.toString q
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   540
  in a ^ " = " ^ s end;
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   541
19049
2103a8e14eaa counter example: avoid vacuous trace;
wenzelm
parents: 18708
diff changeset
   542
fun produce_ex sds =
17496
26535df536ae slight adaptions to library changes
haftmann
parents: 17325
diff changeset
   543
  curry (op ~~) sds
26535df536ae slight adaptions to library changes
haftmann
parents: 17325
diff changeset
   544
  #> map print_atom
26535df536ae slight adaptions to library changes
haftmann
parents: 17325
diff changeset
   545
  #> commas
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   546
  #> curry (op ^) "Counter example (possibly spurious):\n";
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   547
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   548
fun trace_ex (sg, params, atoms, discr, n, hist : history) =
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   549
  case hist of
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   550
    [] => ()
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   551
  | (v, lineqs) :: hist' =>
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   552
    let val frees = map Free params
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   553
        fun s_of_t t = Sign.string_of_term sg (subst_bounds (frees, t))
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   554
        val start = if v = ~1 then (findex0 discr n lineqs, hist')
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   555
                    else (replicate n Rat.zero, hist)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   556
        val ex = SOME (produce_ex ((map s_of_t atoms) ~~ discr) (findex discr start))
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   557
          handle NoEx => NONE
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   558
    in
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   559
      case ex of
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   560
        SOME s => (warning "arith failed - see trace for a counter example"; tracing s)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   561
      | NONE => warning "arith failed"
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   562
    end;
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   563
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   564
(* ------------------------------------------------------------------------- *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   565
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   566
(* Term.typ list -> int list -> Term.term * int -> lineq option *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   567
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   568
fun mknat pTs ixs (atom, i) =
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   569
  if LA_Logic.is_nat (pTs, atom)
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   570
  then let val l = map (fn j => if j=i then 1 else 0) ixs
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   571
       in SOME (Lineq (0, Le, l, Nat i)) end
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   572
  else NONE;
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   573
13186
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   574
(* This code is tricky. It takes a list of premises in the order they occur
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15184
diff changeset
   575
in the subgoal. Numerical premises are coded as SOME(tuple), non-numerical
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15184
diff changeset
   576
ones as NONE. Going through the premises, each numeric one is converted into
13186
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   577
a Lineq. The tricky bit is to convert ~= which is split into two cases < and
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   578
>. Thus split_items returns a list of equation systems. This may blow up if
13186
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   579
there are many ~=, but in practice it does not seem to happen. The really
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   580
tricky bit is to arrange the order of the cases such that they coincide with
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   581
the order in which the cases are in the end generated by the tactic that
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   582
applies the generated refutation thms (see function 'refute_tac').
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   583
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   584
For variables n of type nat, a constraint 0 <= n is added.
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   585
*)
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   586
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   587
(* FIXME: To optimize, the splitting of cases and the search for refutations *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   588
(*        should be intertwined: separate the first (fully split) case,      *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   589
(*        refute it, continue with splitting and refuting.  Terminate with   *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   590
(*        failure as soon as a case could not be refuted; i.e. delay further *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   591
(*        splitting until after a refutation for other cases has been found. *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   592
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   593
(* Theory.theory -> bool -> typ list * term list -> (typ list * (decompT * int) list) list *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   594
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   595
fun split_items sg do_pre (Ts, terms) =
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   596
  let
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   597
(*
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   598
      val _ = trace_msg ("split_items: do_pre is " ^ Bool.toString do_pre ^ "\n" ^
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   599
                         "  Ts    = " ^ string_of_list (Sign.string_of_typ sg) Ts ^ "\n" ^
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   600
                         "  terms = " ^ string_of_list (Sign.string_of_term sg) terms)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   601
*)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   602
      (* splits inequalities '~=' into '<' and '>'; this corresponds to *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   603
      (* 'REPEAT_DETERM (eresolve_tac neqE i)' at the theorem/tactic    *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   604
      (* level                                                          *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   605
      (* decompT option list -> decompT option list list *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   606
      fun elim_neq []              = [[]]
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   607
        | elim_neq (NONE :: ineqs) = map (cons NONE) (elim_neq ineqs)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   608
        | elim_neq (SOME(ineq as (l,i,rel,r,j,d)) :: ineqs) =
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   609
          if rel = "~=" then elim_neq (ineqs @ [SOME (l, i, "<", r, j, d)]) @
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   610
                             elim_neq (ineqs @ [SOME (r, j, "<", l, i, d)])
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   611
          else map (cons (SOME ineq)) (elim_neq ineqs)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   612
      (* int -> decompT option list -> (decompT * int) list *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   613
      fun number_hyps _ []             = []
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   614
        | number_hyps n (NONE::xs)     = number_hyps (n+1) xs
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   615
        | number_hyps n ((SOME x)::xs) = (x, n) :: number_hyps (n+1) xs
13464
c98321b8d638 Fixed two bugs
nipkow
parents: 13186
diff changeset
   616
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   617
      val result = (Ts, terms) |> (* user-defined preprocessing of the subgoal *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   618
                                  (* (typ list * term list) list *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   619
                                  (if do_pre then LA_Data.pre_decomp sg else Library.single)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   620
                               |> (* compute the internal encoding of (in-)equalities *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   621
                                  (* (typ list * decompT option list) list *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   622
                                  map (apsnd (map (LA_Data.decomp sg)))
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   623
                               |> (* splitting of inequalities *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   624
                                  (* (typ list * decompT option list) list list *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   625
                                  map (fn (Ts, items) => map (pair Ts) (elim_neq items))
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   626
                               |> (* combine the list of lists of subgoals into a single list *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   627
                                  (* (typ list * decompT option list) list *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   628
                                  List.concat
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   629
                               |> (* numbering of hypotheses, ignoring irrelevant ones *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   630
                                  (* (typ list * (decompT * int) list) list *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   631
                                  map (apsnd (number_hyps 0))
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   632
(*
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   633
      val _ = trace_msg ("split_items: result has " ^ Int.toString (length result) ^ " subgoal(s)"
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   634
                ^ "\n" ^ (cat_lines o fst) (fold_map (fn (Ts, items) => fn n =>
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   635
                        ("  " ^ Int.toString n ^ ". Ts    = " ^ string_of_list (Sign.string_of_typ sg) Ts ^ "\n" ^
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   636
                         "    items = " ^ string_of_list
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   637
                                            (string_of_pair
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   638
                                              (fn (l, i, rel, r, j, d) =>
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   639
                                                enclose "(" ")" (commas
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   640
                                                  [string_of_list (string_of_pair (Sign.string_of_term sg) Rat.string_of_rat) l,
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   641
                                                   Rat.string_of_rat i,
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   642
                                                   rel,
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   643
                                                   string_of_list (string_of_pair (Sign.string_of_term sg) Rat.string_of_rat) r,
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   644
                                                   Rat.string_of_rat j,
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   645
                                                   Bool.toString d]))
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   646
                                              Int.toString) items, n+1)) result 1))
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   647
*)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   648
  in result end;
13464
c98321b8d638 Fixed two bugs
nipkow
parents: 13186
diff changeset
   649
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   650
(* term list * (decompT * int) -> term list *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   651
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   652
fun add_atoms (ats, ((lhs,_,_,rhs,_,_),_)) =
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   653
    (map fst lhs) union ((map fst rhs) union ats);
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   654
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   655
(* (bool * term) list * (decompT * int) -> (bool * term) list *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   656
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   657
fun add_datoms (dats, ((lhs,_,_,rhs,_,d),_)) =
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   658
    (map (pair d o fst) lhs) union ((map (pair d o fst) rhs) union dats);
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   659
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   660
(* (decompT * int) list -> bool list *)
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   661
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   662
fun discr initems = map fst (Library.foldl add_datoms ([],initems));
13464
c98321b8d638 Fixed two bugs
nipkow
parents: 13186
diff changeset
   663
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   664
(* Theory.theory -> (string * typ) list -> bool -> (typ list * (decompT * int) list) list -> injust list -> injust list option *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   665
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   666
fun refutes sg params show_ex =
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   667
let
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   668
  (* (typ list * (decompT * int) list) list -> injust list -> injust list option *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   669
  fun refute ((Ts, initems)::initemss) js =
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   670
    let val atoms = Library.foldl add_atoms ([], initems)
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   671
        val n = length atoms
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   672
        val mkleq = mklineq n atoms
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   673
        val ixs = 0 upto (n-1)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   674
        val iatoms = atoms ~~ ixs
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   675
        val natlineqs = List.mapPartial (mknat Ts ixs) iatoms
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   676
        val ineqs = map mkleq initems @ natlineqs
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   677
    in case elim (ineqs, []) of
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   678
         Success j =>
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   679
           (trace_msg ("Contradiction! (" ^ Int.toString (length js + 1) ^ ")"); refute initemss (js@[j]))
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   680
       | Failure hist =>
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   681
           (if not show_ex then
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   682
              ()
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   683
            else let
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   684
              (* invent names for bound variables that are new, i.e. in Ts, but not in params *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   685
              (* we assume that Ts still contains (map snd params) as a suffix                *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   686
              val new_count = length Ts - length params - 1
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   687
              val new_names = map Name.bound (0 upto new_count)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   688
              val params'   = (new_names @ map fst params) ~~ Ts
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   689
            in
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   690
              trace_ex (sg, params', atoms, discr initems, n, hist)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   691
            end; NONE)
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   692
    end
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15184
diff changeset
   693
    | refute [] js = SOME js
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   694
in refute end;
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   695
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   696
(* Theory.theory -> (string * Term.typ) list -> bool -> bool -> term list -> injust list option *)
13186
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   697
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   698
fun refute sg params show_ex do_pre terms =
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   699
  refutes sg params show_ex (split_items sg do_pre (map snd params, terms)) [];
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   700
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   701
(* MetaSimplifier.simpset -> int * injust list -> Tactical.tactic *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   702
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   703
fun refute_tac ss (i, justs) =
6074
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   704
  fn state =>
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   705
    let val _ = trace_thm ("refute_tac (on subgoal " ^ Int.toString i ^ ", with " ^ Int.toString (length justs) ^ " justification(s)):") state
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   706
        val sg          = theory_of_thm state
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   707
        val {neqE, ...} = Data.get sg
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   708
        fun just1 j =
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   709
          REPEAT_DETERM (eresolve_tac neqE i) THEN                  (* eliminate inequalities *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   710
            METAHYPS (fn asms => rtac (mkthm (sg, ss) asms j) 1) i  (* use theorems generated from the actual justifications *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   711
    in DETERM (resolve_tac [LA_Logic.notI, LA_Logic.ccontr] i) THEN  (* rewrite "[| A1; ...; An |] ==> B" to "[| A1; ...; An; ~B |] ==> False" *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   712
       DETERM (LA_Data.pre_tac i) THEN                               (* user-defined preprocessing of the subgoal *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   713
       PRIMITIVE (trace_thm "State after pre_tac:") THEN
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   714
       EVERY (map just1 justs)                                       (* prove every resulting subgoal, using its justification *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   715
    end  state;
6074
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   716
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   717
(* ('a -> bool) -> 'a list -> int *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   718
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   719
fun count P xs = length (List.filter P xs);
14510
73ea1234bb23 introduced fast_arith_neq_limit
nipkow
parents: 14386
diff changeset
   720
73ea1234bb23 introduced fast_arith_neq_limit
nipkow
parents: 14386
diff changeset
   721
(* The limit on the number of ~= allowed.
73ea1234bb23 introduced fast_arith_neq_limit
nipkow
parents: 14386
diff changeset
   722
   Because each ~= is split into two cases, this can lead to an explosion.
73ea1234bb23 introduced fast_arith_neq_limit
nipkow
parents: 14386
diff changeset
   723
*)
73ea1234bb23 introduced fast_arith_neq_limit
nipkow
parents: 14386
diff changeset
   724
val fast_arith_neq_limit = ref 9;
73ea1234bb23 introduced fast_arith_neq_limit
nipkow
parents: 14386
diff changeset
   725
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   726
(* Theory.theory -> (string * Term.typ) list -> bool -> bool -> Term.term list -> Term.term -> injust list option *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   727
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   728
fun prove sg params show_ex do_pre Hs concl =
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   729
  let
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   730
    (* append the negated conclusion to 'Hs' -- this corresponds to     *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   731
    (* 'DETERM (resolve_tac [LA_Logic.notI, LA_Logic.ccontr] i)' at the *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   732
    (* theorem/tactic level                                             *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   733
    val Hs' = Hs @ [LA_Logic.neg_prop concl]
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   734
    (* decompT option -> bool *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   735
    fun is_neq NONE                 = false
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   736
      | is_neq (SOME (_,_,r,_,_,_)) = (r = "~=")
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   737
  in
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   738
    trace_msg "prove";
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   739
    if count is_neq (map (LA_Data.decomp sg) Hs')
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   740
      > !fast_arith_neq_limit then (
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   741
      trace_msg ("fast_arith_neq_limit exceeded (current value is " ^ string_of_int (!fast_arith_neq_limit) ^ ")");
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   742
      NONE
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   743
    ) else
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   744
      refute sg params show_ex do_pre Hs'
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   745
  end;
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   746
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   747
(*
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   748
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
   749
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
   750
*)
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   751
fun simpset_lin_arith_tac ss show_ex i st = SUBGOAL (fn (A,_) =>
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   752
  let val params = rev (Logic.strip_params A)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   753
      val Hs     = Logic.strip_assums_hyp A
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   754
      val concl  = Logic.strip_assums_concl A
12932
3bda5306d262 updated debugging output
nipkow
parents: 12311
diff changeset
   755
  in trace_thm ("Trying to refute subgoal " ^ string_of_int i) st;
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   756
     case prove (Thm.sign_of_thm st) params show_ex true Hs concl of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15184
diff changeset
   757
       NONE => (trace_msg "Refutation failed."; no_tac)
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   758
     | SOME js => (trace_msg "Refutation succeeded."; refute_tac ss (i, js))
9420
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   759
  end) i st;
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   760
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   761
fun lin_arith_tac show_ex i st =
17892
62c397c17d18 Simplifier.theory_context;
wenzelm
parents: 17877
diff changeset
   762
  simpset_lin_arith_tac (Simplifier.theory_context (Thm.theory_of_thm st) Simplifier.empty_ss)
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   763
    show_ex i st;
17613
072c21e31b42 Simplifier.inherit_bounds;
wenzelm
parents: 17524
diff changeset
   764
072c21e31b42 Simplifier.inherit_bounds;
wenzelm
parents: 17524
diff changeset
   765
fun cut_lin_arith_tac ss i =
072c21e31b42 Simplifier.inherit_bounds;
wenzelm
parents: 17524
diff changeset
   766
  cut_facts_tac (Simplifier.prems_of_ss ss) i THEN
072c21e31b42 Simplifier.inherit_bounds;
wenzelm
parents: 17524
diff changeset
   767
  simpset_lin_arith_tac ss false i;
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   768
13186
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   769
(** Forward proof from theorems **)
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   770
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   771
(* More tricky code. Needs to arrange the proofs of the multiple cases (due
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   772
to splits of ~= premises) such that it coincides with the order of the cases
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   773
generated by function split_items. *)
13186
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   774
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   775
datatype splittree = Tip of thm list
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   776
                   | Spl of thm * cterm * splittree * cterm * splittree
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   777
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   778
(* the cterm "(ct1 ==> ?R) ==> (ct2 ==> ?R) ==> ?R" is taken to (ct1, ct2) *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   779
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   780
(* Thm.cterm -> Thm.cterm * Thm.cterm *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   781
13186
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   782
fun extract imp =
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   783
let val (Il, r)    = Thm.dest_comb imp
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   784
    val (_, imp1)  = Thm.dest_comb Il
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   785
    val (Ict1, _)  = Thm.dest_comb imp1
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   786
    val (_, ct1)   = Thm.dest_comb Ict1
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   787
    val (Ir, _)    = Thm.dest_comb r
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   788
    val (_, Ict2r) = Thm.dest_comb Ir
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   789
    val (Ict2, _)  = Thm.dest_comb Ict2r
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   790
    val (_, ct2)   = Thm.dest_comb Ict2
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   791
in (ct1, ct2) end;
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   792
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   793
(* Theory.theory -> Thm.thm list -> splittree *)
6074
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   794
15922
7ef183f3fc98 neqE applies even if the type is not one which partakes in linear arithmetic.
nipkow
parents: 15660
diff changeset
   795
fun splitasms sg asms =
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   796
let val {neqE, ...} = Data.get sg
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   797
    fun elim_neq (asms', []) = Tip (rev asms')
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   798
      | elim_neq (asms', asm::asms) =
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   799
      (case get_first (fn th => SOME (asm COMP th) handle THM _ => NONE) neqE of
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   800
        SOME spl =>
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   801
          let val (ct1, ct2) = extract (cprop_of spl)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   802
              val thm1 = assume ct1
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   803
              val thm2 = assume ct2
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   804
          in Spl (spl, ct1, elim_neq (asms', asms@[thm1]), ct2, elim_neq (asms', asms@[thm2]))
15922
7ef183f3fc98 neqE applies even if the type is not one which partakes in linear arithmetic.
nipkow
parents: 15660
diff changeset
   805
          end
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   806
      | NONE => elim_neq (asm::asms', asms))
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   807
in elim_neq ([], asms) end;
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   808
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   809
(* Theory.theory * MetaSimplifier.simpset -> splittree -> injust list -> (Thm.thm, injust list) *)
6074
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   810
17515
830bc15e692c Simplifier.inherit_bounds;
wenzelm
parents: 17496
diff changeset
   811
fun fwdproof ctxt (Tip asms) (j::js) = (mkthm ctxt asms j, js)
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   812
  | fwdproof ctxt (Spl (thm, ct1, tree1, ct2, tree2)) js =
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   813
    let val (thm1, js1) = fwdproof ctxt tree1 js
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   814
        val (thm2, js2) = fwdproof ctxt tree2 js1
13186
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   815
        val thm1' = implies_intr ct1 thm1
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   816
        val thm2' = implies_intr ct2 thm2
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   817
    in (thm2' COMP (thm1' COMP thm), js2) end;
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15184
diff changeset
   818
(* needs handle THM _ => NONE ? *)
13186
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   819
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   820
(* Theory.theory * MetaSimplifier.simpset -> Thm.thm list -> Term.term -> injust list -> bool -> Thm.thm option *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   821
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   822
fun prover (ctxt as (sg, ss)) thms Tconcl js pos =
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   823
let 
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   824
(* vvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvvv *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   825
(* Use this code instead if lin_arith_prover calls prove with do_pre set to true *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   826
    (* There is no "forward version" of 'pre_tac'.  Therefore we combine the     *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   827
    (* available theorems into a single proof state and perform "backward proof" *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   828
    (* using 'refute_tac'.                                                       *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   829
(*
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   830
    val Hs    = map prop_of thms
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   831
    val Prop  = fold (curry Logic.mk_implies) (rev Hs) Tconcl
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   832
    val cProp = cterm_of sg Prop
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   833
    val concl = Goal.init cProp
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   834
                  |> refute_tac ss (1, js)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   835
                  |> Seq.hd
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   836
                  |> Goal.finish
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   837
                  |> fold (fn thA => fn thAB => implies_elim thAB thA) thms
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   838
*)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   839
(* ^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^ *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   840
    val nTconcl       = LA_Logic.neg_prop Tconcl
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   841
    val cnTconcl      = cterm_of sg nTconcl
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   842
    val nTconclthm    = assume cnTconcl
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   843
    val tree          = splitasms sg (thms @ [nTconclthm])
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   844
    val (Falsethm, _) = fwdproof ctxt tree js
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   845
    val contr         = if pos then LA_Logic.ccontr else LA_Logic.notI
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   846
    val concl         = implies_intr cnTconcl Falsethm COMP contr
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   847
in SOME (trace_thm "Proved by lin. arith. prover:"
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   848
          (LA_Logic.mk_Eq concl)) end
13186
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   849
(* in case concl contains ?-var, which makes assume fail: *)
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15184
diff changeset
   850
handle THM _ => NONE;
13186
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   851
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   852
(* PRE: concl is not negated!
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   853
   This assumption is OK because
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   854
   1. lin_arith_prover tries both to prove and disprove concl and
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   855
   2. lin_arith_prover is applied by the simplifier which
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   856
      dives into terms and will thus try the non-negated concl anyway.
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   857
*)
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   858
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   859
(* Theory.theory -> MetaSimplifier.simpset -> Term.term -> Thm.thm option *)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   860
15027
d23887300b96 adapted type of simprocs;
wenzelm
parents: 14821
diff changeset
   861
fun lin_arith_prover sg ss concl =
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   862
let val thms = List.concat (map LA_Logic.atomize (prems_of_ss ss));
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   863
    val Hs = map prop_of thms
6102
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
   864
    val Tconcl = LA_Logic.mk_Trueprop concl
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   865
(*
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   866
    val _ = trace_msg "lin_arith_prover"
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   867
    val _ = map (trace_thm "thms:") thms
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   868
    val _ = trace_msg ("concl:" ^ Sign.string_of_term sg concl)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   869
*)
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   870
in case prove sg [] false false Hs Tconcl of (* concl provable? *)
17515
830bc15e692c Simplifier.inherit_bounds;
wenzelm
parents: 17496
diff changeset
   871
     SOME js => prover (sg, ss) thms Tconcl js true
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15184
diff changeset
   872
   | NONE => let val nTconcl = LA_Logic.neg_prop Tconcl
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 19618
diff changeset
   873
          in case prove sg [] false false Hs nTconcl of (* ~concl provable? *)
17515
830bc15e692c Simplifier.inherit_bounds;
wenzelm
parents: 17496
diff changeset
   874
               SOME js => prover (sg, ss) thms nTconcl js false
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15184
diff changeset
   875
             | NONE => NONE
6079
4f7975c74cdf More arith simplifications.
nipkow
parents: 6074
diff changeset
   876
          end
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   877
end;
6074
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   878
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   879
end;