src/Provers/Arith/fast_lin_arith.ML
author nipkow
Fri, 02 Apr 2004 14:48:31 +0200
changeset 14510 73ea1234bb23
parent 14386 ad1ffcc90162
child 14821 241d2db86fc2
permissions -rw-r--r--
introduced fast_arith_neq_limit
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
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    14
    lin_arith_prover: Sign.sg -> thm list -> term -> thm option
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
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    26
  val conjI:		thm
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 neqE:             thm (* [| m ~= n; m < n ==> P; n < m ==> P |] ==> P *)
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    29
  val notI:             thm (* (P ==> False) ==> ~ P *)
6110
15c2b571225b Simplified interface.
nipkow
parents: 6102
diff changeset
    30
  val not_lessD:        thm (* ~(m < n) ==> n <= m *)
6128
2acc5d36610c More arith refinements.
nipkow
parents: 6110
diff changeset
    31
  val not_leD:          thm (* ~(m <= n) ==> n < m *)
6102
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    32
  val sym:		thm (* x = y ==> y = x *)
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    33
  val mk_Eq: thm -> thm
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
2acc5d36610c More arith refinements.
nipkow
parents: 6110
diff changeset
    38
  val mk_nat_thm: Sign.sg -> 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
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    45
neg_prop(t) = neg if t is wrapped up in Trueprop and
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    46
  nt is the (logically) negated version of t, where the negation
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
6128
2acc5d36610c More arith refinements.
nipkow
parents: 6110
diff changeset
    55
  val decomp:
10691
4ea37fba9c02 towards rtional arithmetic
nipkow
parents: 10575
diff changeset
    56
    Sign.sg -> term -> ((term*rat)list * rat * string * (term*rat)list * rat * bool)option
4ea37fba9c02 towards rtional arithmetic
nipkow
parents: 10575
diff changeset
    57
  val number_of: int * typ -> term
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    58
end;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    59
(*
7551
8e934d1a9ac6 Now distinguishes discrete from non-distrete types.
nipkow
parents: 6128
diff changeset
    60
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
    61
   where Rel is one of "<", "~<", "<=", "~<=" and "=" and
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    62
         p/q is the decomposition of the sum terms x/y into a list
7551
8e934d1a9ac6 Now distinguishes discrete from non-distrete types.
nipkow
parents: 6128
diff changeset
    63
         of summand * multiplicity pairs and a constant summand and
8e934d1a9ac6 Now distinguishes discrete from non-distrete types.
nipkow
parents: 6128
diff changeset
    64
         d indicates if the domain is discrete.
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    65
9420
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
    66
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
    67
   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
    68
   otherwise <= can grow to massive proportions.
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    69
*)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    70
6062
ede9fea461f5 Small mods.
nipkow
parents: 6056
diff changeset
    71
signature FAST_LIN_ARITH =
ede9fea461f5 Small mods.
nipkow
parents: 6056
diff changeset
    72
sig
9420
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
    73
  val setup: (theory -> theory) list
10691
4ea37fba9c02 towards rtional arithmetic
nipkow
parents: 10575
diff changeset
    74
  val map_data: ({add_mono_thms: thm list, mult_mono_thms: (thm*cterm)list, inj_thms: thm list,
10575
c78d26d5c3c1 Now adjusted to mixed terms involving coercions.
nipkow
parents: 9420
diff changeset
    75
                 lessD: thm list, simpset: Simplifier.simpset}
10691
4ea37fba9c02 towards rtional arithmetic
nipkow
parents: 10575
diff changeset
    76
                 -> {add_mono_thms: thm list, mult_mono_thms: (thm*cterm)list, inj_thms: thm list,
10575
c78d26d5c3c1 Now adjusted to mixed terms involving coercions.
nipkow
parents: 9420
diff changeset
    77
                     lessD: thm list, simpset: Simplifier.simpset})
c78d26d5c3c1 Now adjusted to mixed terms involving coercions.
nipkow
parents: 9420
diff changeset
    78
                -> theory -> theory
9073
40d8dfac96b8 tracing flag for arith_tac
paulson
parents: 8263
diff changeset
    79
  val trace           : bool ref
14510
73ea1234bb23 introduced fast_arith_neq_limit
nipkow
parents: 14386
diff changeset
    80
  val fast_arith_neq_limit: int ref
6074
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
    81
  val lin_arith_prover: Sign.sg -> thm list -> term -> thm option
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
    82
  val     lin_arith_tac:     bool -> int -> tactic
6062
ede9fea461f5 Small mods.
nipkow
parents: 6056
diff changeset
    83
  val cut_lin_arith_tac: thm list -> int -> tactic
ede9fea461f5 Small mods.
nipkow
parents: 6056
diff changeset
    84
end;
ede9fea461f5 Small mods.
nipkow
parents: 6056
diff changeset
    85
6102
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    86
functor Fast_Lin_Arith(structure LA_Logic:LIN_ARITH_LOGIC 
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
    87
                       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
    88
struct
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
    89
9420
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
    90
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
    91
(** theory data **)
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
    92
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
    93
(* data kind 'Provers/fast_lin_arith' *)
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
    94
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
    95
structure DataArgs =
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
    96
struct
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
    97
  val name = "Provers/fast_lin_arith";
10691
4ea37fba9c02 towards rtional arithmetic
nipkow
parents: 10575
diff changeset
    98
  type T = {add_mono_thms: thm list, mult_mono_thms: (thm*cterm)list, inj_thms: thm list,
10575
c78d26d5c3c1 Now adjusted to mixed terms involving coercions.
nipkow
parents: 9420
diff changeset
    99
            lessD: thm list, simpset: Simplifier.simpset};
9420
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   100
10691
4ea37fba9c02 towards rtional arithmetic
nipkow
parents: 10575
diff changeset
   101
  val empty = {add_mono_thms = [], mult_mono_thms = [], inj_thms = [],
10575
c78d26d5c3c1 Now adjusted to mixed terms involving coercions.
nipkow
parents: 9420
diff changeset
   102
               lessD = [], simpset = Simplifier.empty_ss};
9420
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   103
  val copy = I;
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   104
  val prep_ext = I;
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   105
10691
4ea37fba9c02 towards rtional arithmetic
nipkow
parents: 10575
diff changeset
   106
  fun merge ({add_mono_thms= add_mono_thms1, mult_mono_thms= mult_mono_thms1, inj_thms= inj_thms1,
10575
c78d26d5c3c1 Now adjusted to mixed terms involving coercions.
nipkow
parents: 9420
diff changeset
   107
              lessD = lessD1, simpset = simpset1},
10691
4ea37fba9c02 towards rtional arithmetic
nipkow
parents: 10575
diff changeset
   108
             {add_mono_thms= add_mono_thms2, mult_mono_thms= mult_mono_thms2, inj_thms= inj_thms2,
10575
c78d26d5c3c1 Now adjusted to mixed terms involving coercions.
nipkow
parents: 9420
diff changeset
   109
              lessD = lessD2, simpset = simpset2}) =
9420
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   110
    {add_mono_thms = Drule.merge_rules (add_mono_thms1, add_mono_thms2),
13105
3d1e7a199bdc use eq_thm_prop instead of slightly inadequate eq_thm;
wenzelm
parents: 12932
diff changeset
   111
     mult_mono_thms = gen_merge_lists' (Drule.eq_thm_prop o pairself fst)
3d1e7a199bdc use eq_thm_prop instead of slightly inadequate eq_thm;
wenzelm
parents: 12932
diff changeset
   112
       mult_mono_thms1 mult_mono_thms2,
10575
c78d26d5c3c1 Now adjusted to mixed terms involving coercions.
nipkow
parents: 9420
diff changeset
   113
     inj_thms = Drule.merge_rules (inj_thms1, inj_thms2),
c78d26d5c3c1 Now adjusted to mixed terms involving coercions.
nipkow
parents: 9420
diff changeset
   114
     lessD = Drule.merge_rules (lessD1, lessD2),
c78d26d5c3c1 Now adjusted to mixed terms involving coercions.
nipkow
parents: 9420
diff changeset
   115
     simpset = Simplifier.merge_ss (simpset1, simpset2)};
9420
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   116
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   117
  fun print _ _ = ();
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   118
end;
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   119
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   120
structure Data = TheoryDataFun(DataArgs);
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   121
val map_data = Data.map;
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   122
val setup = [Data.init];
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   123
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   124
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   125
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   126
(*** A fast decision procedure ***)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   127
(*** Code ported from HOL Light ***)
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   128
(* possible optimizations:
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   129
   use (var,coeff) rep or vector rep  tp save space;
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   130
   treat non-negative atoms separately rather than adding 0 <= atom
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   131
*)
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   132
9073
40d8dfac96b8 tracing flag for arith_tac
paulson
parents: 8263
diff changeset
   133
val trace = ref false;
40d8dfac96b8 tracing flag for arith_tac
paulson
parents: 8263
diff changeset
   134
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   135
datatype lineq_type = Eq | Le | Lt;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   136
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   137
datatype injust = Asm of int
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   138
                | Nat of int (* index of atom *)
6128
2acc5d36610c More arith refinements.
nipkow
parents: 6110
diff changeset
   139
                | LessD of injust
2acc5d36610c More arith refinements.
nipkow
parents: 6110
diff changeset
   140
                | NotLessD of injust
2acc5d36610c More arith refinements.
nipkow
parents: 6110
diff changeset
   141
                | NotLeD of injust
7551
8e934d1a9ac6 Now distinguishes discrete from non-distrete types.
nipkow
parents: 6128
diff changeset
   142
                | NotLeDD of injust
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   143
                | Multiplied of int * injust
10691
4ea37fba9c02 towards rtional arithmetic
nipkow
parents: 10575
diff changeset
   144
                | Multiplied2 of int * injust
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   145
                | Added of injust * injust;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   146
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   147
datatype lineq = Lineq of int * lineq_type * int list * injust;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   148
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   149
fun el 0 (h::_) = h
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   150
  | el n (_::t) = el (n - 1) t
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   151
  | el _ _  = sys_error "el";
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   152
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   153
(* ------------------------------------------------------------------------- *)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   154
(* Finding a (counter) example from the trace of a failed elimination        *)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   155
(* ------------------------------------------------------------------------- *)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   156
(* Examples are represented as rational numbers,                             *)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   157
(* Dont blame John Harrison for this code - it is entirely mine. TN          *)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   158
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   159
exception NoEx;
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   160
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   161
(* 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
   162
   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
   163
   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
   164
*)
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   165
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   166
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
   167
  | 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
   168
  | elim_eqns(ineqs,Lineq(i,Lt,cs,_)) = (i,false,cs)::ineqs;
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   169
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   170
val rat0 = rat_of_int 0;
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   171
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   172
(* PRE: ex[v] must be 0! *)
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   173
fun eval (ex:rat list) v (a:int,le,cs:int list) =
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   174
  let val rs = map rat_of_int cs
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   175
      val rsum = foldl ratadd (rat0,map ratmul (rs ~~ ex))
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   176
  in (ratmul(ratadd(rat_of_int a,ratneg rsum), ratinv(el v rs)), le) end;
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   177
(* 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
   178
   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
   179
*)
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   180
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   181
fun ratge0 r = fst(rep_rat r) >= 0;
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   182
fun ratle(r,s) = ratge0(ratadd(s,ratneg r));
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   183
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   184
fun ratrelmin2(x as (r,ler),y as (s,les)) =
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   185
  if r=s then (r, (not ler) andalso (not les)) else if ratle(r,s) then x else y;
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   186
fun ratrelmax2(x as (r,ler),y as (s,les)) =
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   187
  if r=s then (r,ler andalso les) else if ratle(r,s) then y else x;
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
val ratrelmin = foldr1 ratrelmin2;
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   190
val ratrelmax = foldr1 ratrelmax2;
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   191
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   192
fun ratroundup r = let val (p,q) = rep_rat r
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   193
                   in if q=1 then r else rat_of_int((p div q) + 1) end
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   194
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   195
fun ratrounddown r = let val (p,q) = rep_rat r
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   196
                     in if q=1 then r else rat_of_int((p div q) - 1) end
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   197
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   198
fun ratexact up (r,exact) =
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   199
  if exact then r else
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   200
  let val (p,q) = rep_rat r
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   201
      val nth = ratinv(rat_of_int q)
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   202
  in ratadd(r,if up then nth else ratneg nth) end;
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   203
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   204
fun ratmiddle(r,s) = ratmul(ratadd(r,s),ratinv(rat_of_int 2));
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   205
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   206
fun choose2 d ((lb,exactl),(ub,exactu)) =
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   207
  if ratle(lb,rat0) andalso (lb <> rat0 orelse exactl) andalso
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   208
     ratle(rat0,ub) andalso (ub <> rat0 orelse exactu)
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   209
  then rat0 else
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   210
  if not d
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   211
  then (if ratge0 lb
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   212
        then if exactl then lb else ratmiddle(lb,ub)
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   213
        else if exactu then ub else ratmiddle(lb,ub))
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   214
  else (* discrete domain, both bounds must be exact *)
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   215
  if ratge0 lb then let val lb' = ratroundup lb
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   216
                    in if ratle(lb',ub) then lb' else raise NoEx end
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   217
               else let val ub' = ratrounddown ub
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   218
                    in if ratle(lb,ub') then ub' else raise NoEx end;
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   219
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   220
fun findex1 discr (ex,(v,lineqs)) =
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   221
  let val nz = filter (fn (Lineq(_,_,cs,_)) => el v cs <> 0) lineqs;
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   222
      val ineqs = foldl elim_eqns ([],nz)
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   223
      val (ge,le) = partition (fn (_,_,cs) => el v cs > 0) ineqs
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   224
      val lb = ratrelmax(map (eval ex v) ge)
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   225
      val ub = ratrelmin(map (eval ex v) le)
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   226
  in nth_update (choose2 (nth_elem(v,discr)) (lb,ub)) (v,ex) end;
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   227
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   228
fun findex discr = foldl (findex1 discr);
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   229
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   230
fun elim1 v x =
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   231
  map (fn (a,le,bs) => (ratadd(a,ratneg(ratmul(el v bs,x))), le,
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   232
                        nth_update rat0 (v,bs)));
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   233
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   234
fun single_var v (_,_,cs) = (filter_out (equal rat0) cs = [el v cs]);
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   235
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   236
(* The base case:
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   237
   all variables occur only with positive or only with negative coefficients *)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   238
fun pick_vars discr (ineqs,ex) =
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   239
  let val nz = filter_out (fn (_,_,cs) => forall (equal rat0) cs) ineqs
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   240
  in case nz of [] => ex
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   241
     | (_,_,cs) :: _ =>
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   242
       let val v = find_index (not o equal rat0) cs
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   243
           val d = nth_elem(v,discr)
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   244
           val pos = ratge0(el v cs)
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   245
           val sv = filter (single_var v) nz
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   246
           val minmax =
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   247
             if pos then if d then ratroundup o fst o ratrelmax
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   248
                         else ratexact true o ratrelmax
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   249
                    else if d then ratrounddown o fst o ratrelmin
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   250
                         else ratexact false o ratrelmin
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   251
           val bnds = map (fn (a,le,bs) => (ratmul(a,ratinv(el v bs)),le)) sv
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   252
           val x = minmax((rat0,if pos then true else false)::bnds)
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   253
           val ineqs' = elim1 v x nz
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   254
           val ex' = nth_update x (v,ex)
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   255
       in pick_vars discr (ineqs',ex') end
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   256
  end;
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   257
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   258
fun findex0 discr n lineqs =
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   259
  let val ineqs = foldl elim_eqns ([],lineqs)
14372
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   260
      val rineqs = map (fn (a,le,cs) => (rat_of_int a, le, map rat_of_int cs))
51ddf8963c95 Finally fixed the counterexample finder. Can now deal with < on real.
nipkow
parents: 14360
diff changeset
   261
                       ineqs
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   262
  in pick_vars discr (rineqs,replicate n rat0) end;
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   263
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   264
(* ------------------------------------------------------------------------- *)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   265
(* End of counter example finder. The actual decision procedure starts here. *)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   266
(* ------------------------------------------------------------------------- *)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   267
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   268
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   269
(* 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
   270
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   271
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   272
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
   273
  | find_add_type(x,Eq) = x
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(Lt,_) = Lt
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   276
  | find_add_type(Le,Le) = Le;
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
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   279
(* Multiply out an (in)equation.                                             *)
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
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   282
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
   283
  if n = 1 then i
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 = Lt then sys_error "multiply_ineq"
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   285
  else if n < 0 andalso (ty=Le orelse ty=Lt) then sys_error "multiply_ineq"
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   286
  else Lineq(n * k,ty,map (apl(n,op * )) l,Multiplied(n,just));
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   287
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   288
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   289
(* Add together (in)equations.                                               *)
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
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   292
fun add_ineq (i1 as Lineq(k1,ty1,l1,just1)) (i2 as Lineq(k2,ty2,l2,just2)) =
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   293
  let val l = map2 (op +) (l1,l2)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   294
  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
   295
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   296
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   297
(* 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
   298
(* 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
   299
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   300
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   301
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
   302
  let val c1 = el v l1 and c2 = el v l2
10691
4ea37fba9c02 towards rtional arithmetic
nipkow
parents: 10575
diff changeset
   303
      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
   304
      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
   305
      val (n1,n2) =
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   306
        if (c1 >= 0) = (c2 >= 0)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   307
        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
   308
             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
   309
                  else sys_error "elim_var"
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   310
        else (m1,m2)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   311
      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
   312
                    then (~n1,~n2) else (n1,n2)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   313
  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
   314
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   315
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   316
(* The main refutation-finding code.                                         *)
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
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   319
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
   320
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   321
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
   322
  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
   323
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   324
fun calc_blowup l =
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   325
  let val (p,n) = partition (apl(0,op<)) (filter (apl(0,op<>)) l)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   326
  in (length p) * (length n) end;
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
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   329
(* Main elimination code:                                                    *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   330
(*                                                                           *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   331
(* (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
   332
(*                                                                           *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   333
(* (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
   334
(* 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
   335
(*                                                                           *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   336
(* (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
   337
(* 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
   338
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   339
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   340
fun allpairs f xs ys =
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   341
  flat(map (fn x => map (fn y => f x y) ys) xs);
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   342
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   343
fun extract_first p =
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   344
  let fun extract xs (y::ys) = if p y then (Some y,xs@ys)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   345
                               else extract (y::xs) ys
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   346
        | extract xs []      = (None,xs)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   347
  in extract [] end;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   348
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   349
fun print_ineqs ineqs =
9073
40d8dfac96b8 tracing flag for arith_tac
paulson
parents: 8263
diff changeset
   350
  if !trace then
12262
11ff5f47df6e use tracing function for trace output;
wenzelm
parents: 12109
diff changeset
   351
     tracing(cat_lines(""::map (fn Lineq(c,t,l,_) =>
9073
40d8dfac96b8 tracing flag for arith_tac
paulson
parents: 8263
diff changeset
   352
       string_of_int c ^
40d8dfac96b8 tracing flag for arith_tac
paulson
parents: 8263
diff changeset
   353
       (case t of Eq => " =  " | Lt=> " <  " | Le => " <= ") ^
40d8dfac96b8 tracing flag for arith_tac
paulson
parents: 8263
diff changeset
   354
       commas(map string_of_int l)) ineqs))
40d8dfac96b8 tracing flag for arith_tac
paulson
parents: 8263
diff changeset
   355
  else ();
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   356
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   357
type history = (int * lineq list) list;
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   358
datatype result = Success of injust | Failure of history;
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   359
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   360
fun elim(ineqs,hist) =
9073
40d8dfac96b8 tracing flag for arith_tac
paulson
parents: 8263
diff changeset
   361
  let val dummy = print_ineqs ineqs;
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   362
      val (triv,nontriv) = partition is_trivial ineqs in
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   363
  if not(null triv)
13186
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   364
  then case Library.find_first is_answer triv of
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   365
         None => elim(nontriv,hist)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   366
       | Some(Lineq(_,_,_,j)) => Success j
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   367
  else
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   368
  if null nontriv then Failure(hist)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   369
  else
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   370
  let val (eqs,noneqs) = partition (fn (Lineq(_,ty,_,_)) => ty=Eq) nontriv in
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   371
  if not(null eqs) then
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   372
     let val clist = foldl (fn (cs,Lineq(_,_,l,_)) => l union cs) ([],eqs)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   373
         val sclist = sort (fn (x,y) => int_ord(abs(x),abs(y)))
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   374
                           (filter (fn i => i<>0) clist)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   375
         val c = hd sclist
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   376
         val (Some(eq as Lineq(_,_,ceq,_)),othereqs) =
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   377
               extract_first (fn Lineq(_,_,l,_) => c mem l) eqs
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   378
         val v = find_index_eq c ceq
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   379
         val (ioth,roth) = partition (fn (Lineq(_,_,l,_)) => el v l = 0)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   380
                                     (othereqs @ noneqs)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   381
         val others = map (elim_var v eq) roth @ ioth
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   382
     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
   383
  else
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   384
  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
   385
      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
   386
      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
   387
      val blows = map calc_blowup coeffs
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   388
      val iblows = blows ~~ numlist
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   389
      val nziblows = filter (fn (i,_) => i<>0) iblows
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   390
  in if null nziblows then Failure((~1,nontriv)::hist)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   391
     else
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   392
     let val (c,v) = hd(sort (fn (x,y) => int_ord(fst(x),fst(y))) nziblows)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   393
         val (no,yes) = partition (fn (Lineq(_,_,l,_)) => el v l = 0) ineqs
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   394
         val (pos,neg) = partition(fn (Lineq(_,_,l,_)) => el v l > 0) yes
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   395
     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
   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
  end;
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
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   401
(* Translate back a proof.                                                   *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   402
(* ------------------------------------------------------------------------- *)
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   403
9073
40d8dfac96b8 tracing flag for arith_tac
paulson
parents: 8263
diff changeset
   404
fun trace_thm msg th = 
12262
11ff5f47df6e use tracing function for trace output;
wenzelm
parents: 12109
diff changeset
   405
    if !trace then (tracing msg; tracing (Display.string_of_thm th); th) else th;
9073
40d8dfac96b8 tracing flag for arith_tac
paulson
parents: 8263
diff changeset
   406
40d8dfac96b8 tracing flag for arith_tac
paulson
parents: 8263
diff changeset
   407
fun trace_msg msg = 
12262
11ff5f47df6e use tracing function for trace output;
wenzelm
parents: 12109
diff changeset
   408
    if !trace then tracing msg else ();
9073
40d8dfac96b8 tracing flag for arith_tac
paulson
parents: 8263
diff changeset
   409
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   410
(* FIXME OPTIMIZE!!!! (partly done already)
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   411
   Addition/Multiplication need i*t representation rather than t+t+...
10691
4ea37fba9c02 towards rtional arithmetic
nipkow
parents: 10575
diff changeset
   412
   Get rid of Mulitplied(2). For Nat LA_Data.number_of should return Suc^n
4ea37fba9c02 towards rtional arithmetic
nipkow
parents: 10575
diff changeset
   413
   because Numerals are not known early enough.
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   414
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   415
Simplification may detect a contradiction 'prematurely' due to type
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   416
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
   417
with 0 <= n.
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   418
*)
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   419
local
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   420
 exception FalseE of thm
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   421
in
6074
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   422
fun mkthm sg asms just =
10691
4ea37fba9c02 towards rtional arithmetic
nipkow
parents: 10575
diff changeset
   423
  let val {add_mono_thms, mult_mono_thms, inj_thms, lessD, simpset} = Data.get_sg sg;
9420
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   424
      val atoms = foldl (fn (ats,(lhs,_,_,rhs,_,_)) =>
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   425
                            map fst lhs  union  (map fst rhs  union  ats))
13464
c98321b8d638 Fixed two bugs
nipkow
parents: 13186
diff changeset
   426
                        ([], mapfilter (fn thm => if Thm.no_prems thm
c98321b8d638 Fixed two bugs
nipkow
parents: 13186
diff changeset
   427
                                        then LA_Data.decomp sg (concl_of thm)
c98321b8d638 Fixed two bugs
nipkow
parents: 13186
diff changeset
   428
                                        else None) asms)
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   429
10575
c78d26d5c3c1 Now adjusted to mixed terms involving coercions.
nipkow
parents: 9420
diff changeset
   430
      fun add2 thm1 thm2 =
6102
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
   431
        let val conj = thm1 RS (thm2 RS LA_Logic.conjI)
10575
c78d26d5c3c1 Now adjusted to mixed terms involving coercions.
nipkow
parents: 9420
diff changeset
   432
        in get_first (fn th => Some(conj RS th) handle _ => None) add_mono_thms
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   433
        end;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   434
10575
c78d26d5c3c1 Now adjusted to mixed terms involving coercions.
nipkow
parents: 9420
diff changeset
   435
      fun try_add [] _ = None
c78d26d5c3c1 Now adjusted to mixed terms involving coercions.
nipkow
parents: 9420
diff changeset
   436
        | try_add (thm1::thm1s) thm2 = case add2 thm1 thm2 of
c78d26d5c3c1 Now adjusted to mixed terms involving coercions.
nipkow
parents: 9420
diff changeset
   437
             None => try_add thm1s thm2 | some => some;
c78d26d5c3c1 Now adjusted to mixed terms involving coercions.
nipkow
parents: 9420
diff changeset
   438
c78d26d5c3c1 Now adjusted to mixed terms involving coercions.
nipkow
parents: 9420
diff changeset
   439
      fun addthms thm1 thm2 =
c78d26d5c3c1 Now adjusted to mixed terms involving coercions.
nipkow
parents: 9420
diff changeset
   440
        case add2 thm1 thm2 of
c78d26d5c3c1 Now adjusted to mixed terms involving coercions.
nipkow
parents: 9420
diff changeset
   441
          None => (case try_add ([thm1] RL inj_thms) thm2 of
14360
e654599b114e Added an exception handler and error msg.
nipkow
parents: 13561
diff changeset
   442
                     None => ( the(try_add ([thm2] RL inj_thms) thm1)
e654599b114e Added an exception handler and error msg.
nipkow
parents: 13561
diff changeset
   443
                               handle OPTION =>
e654599b114e Added an exception handler and error msg.
nipkow
parents: 13561
diff changeset
   444
                               (trace_thm "" thm1; trace_thm "" thm2;
e654599b114e Added an exception handler and error msg.
nipkow
parents: 13561
diff changeset
   445
                                sys_error "Lin.arith. failed to add thms")
e654599b114e Added an exception handler and error msg.
nipkow
parents: 13561
diff changeset
   446
                             )
10575
c78d26d5c3c1 Now adjusted to mixed terms involving coercions.
nipkow
parents: 9420
diff changeset
   447
                   | Some thm => thm)
c78d26d5c3c1 Now adjusted to mixed terms involving coercions.
nipkow
parents: 9420
diff changeset
   448
        | Some thm => thm;
c78d26d5c3c1 Now adjusted to mixed terms involving coercions.
nipkow
parents: 9420
diff changeset
   449
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   450
      fun multn(n,thm) =
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   451
        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
   452
        in if n < 0 then mul(~n,thm) RS LA_Logic.sym else mul(n,thm) end;
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   453
10691
4ea37fba9c02 towards rtional arithmetic
nipkow
parents: 10575
diff changeset
   454
      fun multn2(n,thm) =
4ea37fba9c02 towards rtional arithmetic
nipkow
parents: 10575
diff changeset
   455
        let val Some(mth,cv) =
4ea37fba9c02 towards rtional arithmetic
nipkow
parents: 10575
diff changeset
   456
              get_first (fn (th,cv) => Some(thm RS th,cv) handle _ => None) mult_mono_thms
4ea37fba9c02 towards rtional arithmetic
nipkow
parents: 10575
diff changeset
   457
            val ct = cterm_of sg (LA_Data.number_of(n,#T(rep_cterm cv)))
4ea37fba9c02 towards rtional arithmetic
nipkow
parents: 10575
diff changeset
   458
        in instantiate ([],[(cv,ct)]) mth end
4ea37fba9c02 towards rtional arithmetic
nipkow
parents: 10575
diff changeset
   459
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   460
      fun simp thm =
12932
3bda5306d262 updated debugging output
nipkow
parents: 12311
diff changeset
   461
        let val thm' = trace_thm "Simplified:" (full_simplify simpset thm)
6102
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
   462
        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
   463
9073
40d8dfac96b8 tracing flag for arith_tac
paulson
parents: 8263
diff changeset
   464
      fun mk(Asm i) = trace_thm "Asm" (nth_elem(i,asms))
13464
c98321b8d638 Fixed two bugs
nipkow
parents: 13186
diff changeset
   465
        | mk(Nat i) = (trace_msg "Nat"; LA_Logic.mk_nat_thm sg (nth_elem(i,atoms)))
9420
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   466
        | mk(LessD(j)) = trace_thm "L" (hd([mk j] RL lessD))
9073
40d8dfac96b8 tracing flag for arith_tac
paulson
parents: 8263
diff changeset
   467
        | mk(NotLeD(j)) = trace_thm "NLe" (mk j RS LA_Logic.not_leD)
9420
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   468
        | mk(NotLeDD(j)) = trace_thm "NLeD" (hd([mk j RS LA_Logic.not_leD] RL lessD))
9073
40d8dfac96b8 tracing flag for arith_tac
paulson
parents: 8263
diff changeset
   469
        | mk(NotLessD(j)) = trace_thm "NL" (mk j RS LA_Logic.not_lessD)
40d8dfac96b8 tracing flag for arith_tac
paulson
parents: 8263
diff changeset
   470
        | mk(Added(j1,j2)) = simp (trace_thm "+" (addthms (mk j1) (mk j2)))
10717
c09d4ebfec83 rational arithemtic
nipkow
parents: 10691
diff changeset
   471
        | mk(Multiplied(n,j)) = (trace_msg("*"^string_of_int n); trace_thm "*" (multn(n,mk j)))
c09d4ebfec83 rational arithemtic
nipkow
parents: 10691
diff changeset
   472
        | mk(Multiplied2(n,j)) = simp (trace_msg("**"^string_of_int 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
   473
9073
40d8dfac96b8 tracing flag for arith_tac
paulson
parents: 8263
diff changeset
   474
  in trace_msg "mkthm";
12932
3bda5306d262 updated debugging output
nipkow
parents: 12311
diff changeset
   475
     let val thm = trace_thm "Final thm:" (mk just)
3bda5306d262 updated debugging output
nipkow
parents: 12311
diff changeset
   476
     in let val fls = simplify simpset thm
13186
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   477
        in trace_thm "After simplification:" fls;
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   478
           if LA_Logic.is_False fls then fls
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   479
           else
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   480
            (tracing "Assumptions:"; seq print_thm asms;
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   481
             tracing "Proved:"; print_thm fls;
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   482
             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
   483
                     \Please inform Tobias Nipkow (nipkow@in.tum.de).";
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   484
             fls)
12932
3bda5306d262 updated debugging output
nipkow
parents: 12311
diff changeset
   485
        end
3bda5306d262 updated debugging output
nipkow
parents: 12311
diff changeset
   486
     end handle FalseE thm => (trace_thm "False reached early:" thm; thm)
3bda5306d262 updated debugging output
nipkow
parents: 12311
diff changeset
   487
  end
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   488
end;
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   489
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   490
fun coeff poly atom = case assoc(poly,atom) of None => 0 | Some i => i;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   491
10691
4ea37fba9c02 towards rtional arithmetic
nipkow
parents: 10575
diff changeset
   492
fun lcms is = foldl lcm (1,is);
4ea37fba9c02 towards rtional arithmetic
nipkow
parents: 10575
diff changeset
   493
4ea37fba9c02 towards rtional arithmetic
nipkow
parents: 10575
diff changeset
   494
fun integ(rlhs,r,rel,rrhs,s,d) =
4ea37fba9c02 towards rtional arithmetic
nipkow
parents: 10575
diff changeset
   495
let val (rn,rd) = rep_rat r and (sn,sd) = rep_rat s
4ea37fba9c02 towards rtional arithmetic
nipkow
parents: 10575
diff changeset
   496
    val m = lcms(map (abs o snd o rep_rat) (r :: s :: map snd rlhs @ map snd rrhs))
4ea37fba9c02 towards rtional arithmetic
nipkow
parents: 10575
diff changeset
   497
    fun mult(t,r) = let val (i,j) = rep_rat r in (t,i * (m div j)) end
12932
3bda5306d262 updated debugging output
nipkow
parents: 12311
diff changeset
   498
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
   499
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   500
fun mklineq n atoms =
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   501
  fn (item,k) =>
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   502
  let val (m,(lhs,i,rel,rhs,j,discrete)) = integ item
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   503
      val lhsa = map (coeff lhs) atoms
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   504
      and rhsa = map (coeff rhs) atoms
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   505
      val diff = map2 (op -) (rhsa,lhsa)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   506
      val c = i-j
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   507
      val just = Asm k
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   508
      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
   509
  in case rel of
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   510
      "<="   => lineq(c,Le,diff,just)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   511
     | "~<=" => if discrete
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   512
                then lineq(1-c,Le,map (op ~) diff,NotLeDD(just))
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   513
                else lineq(~c,Lt,map (op ~) diff,NotLeD(just))
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   514
     | "<"   => if discrete
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   515
                then lineq(c+1,Le,diff,LessD(just))
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   516
                else lineq(c,Lt,diff,just)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   517
     | "~<"  => lineq(~c,Le,map (op~) diff,NotLessD(just))
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   518
     | "="   => lineq(c,Eq,diff,just)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   519
     | _     => sys_error("mklineq" ^ rel)   
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   520
  end;
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   521
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   522
(* ------------------------------------------------------------------------- *)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   523
(* Print (counter) example                                                   *)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   524
(* ------------------------------------------------------------------------- *)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   525
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   526
fun print_atom((a,d),r) =
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   527
  let val (p,q) = rep_rat r
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   528
      val s = if d then string_of_int p else
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   529
              if p = 0 then "0"
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   530
              else string_of_int p ^ "/" ^ string_of_int q
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   531
  in a ^ " = " ^ s end;
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   532
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   533
fun print_ex sds =
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   534
  tracing o
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   535
  apl("Counter example:\n",op ^) o
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   536
  commas o
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   537
  map print_atom o
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   538
  apl(sds, op ~~);
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   539
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   540
fun trace_ex(sg,params,atoms,discr,n,hist:history) =
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   541
  if null hist then ()
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   542
  else let val frees = map Free params;
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   543
           fun s_of_t t = Sign.string_of_term sg (subst_bounds(frees,t));
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   544
           val (v,lineqs) :: hist' = hist
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   545
           val start = if v = ~1 then (findex0 discr n lineqs,hist')
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   546
                       else (replicate n rat0,hist)
13516
13a6103b9ac4 for cancelling div + mod.
nipkow
parents: 13498
diff changeset
   547
       in warning "arith failed - see trace for a counter example";
13a6103b9ac4 for cancelling div + mod.
nipkow
parents: 13498
diff changeset
   548
          print_ex ((map s_of_t atoms)~~discr) (findex discr start)
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   549
          handle NoEx =>
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   550
  (tracing "The decision procedure failed to prove your proposition\n\
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   551
           \but could not construct a counter example either.\n\
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   552
           \Probably the proposition is true but cannot be proved\n\
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   553
           \by the incomplete decision procedure.")
14386
ad1ffcc90162 Removed dangling exception handler
nipkow
parents: 14372
diff changeset
   554
       end;
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   555
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   556
fun mknat pTs ixs (atom,i) =
6128
2acc5d36610c More arith refinements.
nipkow
parents: 6110
diff changeset
   557
  if LA_Logic.is_nat(pTs,atom)
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   558
  then let val l = map (fn j => if j=i then 1 else 0) ixs
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   559
       in Some(Lineq(0,Le,l,Nat(i))) end
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   560
  else None
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   561
13186
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   562
(* This code is tricky. It takes a list of premises in the order they occur
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   563
in the subgoal. Numerical premises are coded as Some(tuple), non-numerical
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   564
ones as None. Going through the premises, each numeric one is converted into
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   565
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
   566
>. 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
   567
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
   568
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
   569
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
   570
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
   571
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   572
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
   573
*)
13464
c98321b8d638 Fixed two bugs
nipkow
parents: 13186
diff changeset
   574
fun split_items(items) =
c98321b8d638 Fixed two bugs
nipkow
parents: 13186
diff changeset
   575
  let fun elim_neq front _ [] = [rev front]
c98321b8d638 Fixed two bugs
nipkow
parents: 13186
diff changeset
   576
        | elim_neq front n (None::ineqs) = elim_neq front (n+1) ineqs
c98321b8d638 Fixed two bugs
nipkow
parents: 13186
diff changeset
   577
        | elim_neq front n (Some(ineq as (l,i,rel,r,j,d))::ineqs) =
c98321b8d638 Fixed two bugs
nipkow
parents: 13186
diff changeset
   578
          if rel = "~=" then elim_neq front n (ineqs @ [Some(l,i,"<",r,j,d)]) @
c98321b8d638 Fixed two bugs
nipkow
parents: 13186
diff changeset
   579
                             elim_neq front n (ineqs @ [Some(r,j,"<",l,i,d)])
c98321b8d638 Fixed two bugs
nipkow
parents: 13186
diff changeset
   580
          else elim_neq ((ineq,n) :: front) (n+1) ineqs
c98321b8d638 Fixed two bugs
nipkow
parents: 13186
diff changeset
   581
  in elim_neq [] 0 items end;
c98321b8d638 Fixed two bugs
nipkow
parents: 13186
diff changeset
   582
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   583
fun add_atoms(ats,((lhs,_,_,rhs,_,_),_)) =
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   584
    (map fst lhs) union ((map fst rhs) union ats)
13464
c98321b8d638 Fixed two bugs
nipkow
parents: 13186
diff changeset
   585
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   586
fun add_datoms(dats,((lhs,_,_,rhs,_,d),_)) =
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   587
    (map (pair d o fst) lhs) union ((map (pair d o fst) rhs) union dats)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   588
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   589
fun discr initems = map fst (foldl add_datoms ([],initems));
13464
c98321b8d638 Fixed two bugs
nipkow
parents: 13186
diff changeset
   590
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   591
fun refutes sg (pTs,params) ex =
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   592
let
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   593
  fun refute (initems::initemss) js =
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   594
    let val atoms = foldl add_atoms ([],initems)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   595
        val n = length atoms
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   596
        val mkleq = mklineq n atoms
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   597
        val ixs = 0 upto (n-1)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   598
        val iatoms = atoms ~~ ixs
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   599
        val natlineqs = mapfilter (mknat pTs ixs) iatoms
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   600
        val ineqs = map mkleq initems @ natlineqs
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   601
    in case elim(ineqs,[]) of
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   602
         Success(j) =>
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   603
           (trace_msg "Contradiction!"; refute initemss (js@[j]))
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   604
       | Failure(hist) =>
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   605
           (if not ex then ()
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   606
            else trace_ex(sg,params,atoms,discr initems,n,hist);
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   607
            None)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   608
    end
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   609
    | refute [] js = Some js
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   610
in refute end;
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   611
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   612
fun refute sg ps ex items = refutes sg ps ex (split_items items) [];
13186
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   613
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   614
fun refute_tac(i,justs) =
6074
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   615
  fn state =>
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   616
    let val sg = #sign(rep_thm state)
13186
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   617
        fun just1 j = REPEAT_DETERM(etac LA_Logic.neqE i) THEN
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   618
                      METAHYPS (fn asms => rtac (mkthm sg asms j) 1) i
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   619
    in DETERM(resolve_tac [LA_Logic.notI,LA_Logic.ccontr] i) THEN
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   620
       EVERY(map just1 justs)
6074
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   621
    end
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   622
    state;
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   623
14510
73ea1234bb23 introduced fast_arith_neq_limit
nipkow
parents: 14386
diff changeset
   624
fun count P xs = length(filter P xs);
73ea1234bb23 introduced fast_arith_neq_limit
nipkow
parents: 14386
diff changeset
   625
73ea1234bb23 introduced fast_arith_neq_limit
nipkow
parents: 14386
diff changeset
   626
(* The limit on the number of ~= allowed.
73ea1234bb23 introduced fast_arith_neq_limit
nipkow
parents: 14386
diff changeset
   627
   Because each ~= is split into two cases, this can lead to an explosion.
73ea1234bb23 introduced fast_arith_neq_limit
nipkow
parents: 14386
diff changeset
   628
*)
73ea1234bb23 introduced fast_arith_neq_limit
nipkow
parents: 14386
diff changeset
   629
val fast_arith_neq_limit = ref 9;
73ea1234bb23 introduced fast_arith_neq_limit
nipkow
parents: 14386
diff changeset
   630
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   631
fun prove sg ps ex Hs concl =
13186
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   632
let val Hitems = map (LA_Data.decomp sg) Hs
14510
73ea1234bb23 introduced fast_arith_neq_limit
nipkow
parents: 14386
diff changeset
   633
in if count (fn None => false | Some(_,_,r,_,_,_) => r = "~=") Hitems
73ea1234bb23 introduced fast_arith_neq_limit
nipkow
parents: 14386
diff changeset
   634
      > !fast_arith_neq_limit then None
73ea1234bb23 introduced fast_arith_neq_limit
nipkow
parents: 14386
diff changeset
   635
   else
73ea1234bb23 introduced fast_arith_neq_limit
nipkow
parents: 14386
diff changeset
   636
   case LA_Data.decomp sg concl of
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   637
     None => refute sg ps ex (Hitems@[None])
7551
8e934d1a9ac6 Now distinguishes discrete from non-distrete types.
nipkow
parents: 6128
diff changeset
   638
   | Some(citem as (r,i,rel,l,j,d)) =>
13186
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   639
       let val neg::rel0 = explode rel
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   640
           val nrel = if neg = "~" then implode rel0 else "~"^rel
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   641
       in refute sg ps ex (Hitems @ [Some(r,i,nrel,l,j,d)]) end
6074
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   642
end;
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   643
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   644
(*
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   645
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
   646
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
   647
*)
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   648
fun lin_arith_tac ex i st = SUBGOAL (fn (A,_) =>
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   649
  let val params = rev(Logic.strip_params A)
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   650
      val pTs = map snd params
6056
b21813d1b701 Version 1 of linear arithmetic for nat.
nipkow
parents: 5982
diff changeset
   651
      val Hs = Logic.strip_assums_hyp A
6074
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   652
      val concl = Logic.strip_assums_concl A
12932
3bda5306d262 updated debugging output
nipkow
parents: 12311
diff changeset
   653
  in trace_thm ("Trying to refute subgoal " ^ string_of_int i) st;
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   654
     case prove (Thm.sign_of_thm st) (pTs,params) ex Hs concl of
13464
c98321b8d638 Fixed two bugs
nipkow
parents: 13186
diff changeset
   655
       None => (trace_msg "Refutation failed."; no_tac)
c98321b8d638 Fixed two bugs
nipkow
parents: 13186
diff changeset
   656
     | Some js => (trace_msg "Refutation succeeded."; refute_tac(i,js))
9420
d4e9f60fe25a avoid global references;
wenzelm
parents: 9073
diff changeset
   657
  end) i st;
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   658
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   659
fun cut_lin_arith_tac thms i = cut_facts_tac thms i THEN lin_arith_tac false i;
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   660
13186
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   661
(** Forward proof from theorems **)
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   662
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   663
(* 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
   664
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
   665
generated by function split_items. *)
13186
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   666
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   667
datatype splittree = Tip of thm list
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   668
                   | Spl of thm * cterm * splittree * cterm * splittree
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   669
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   670
fun extract imp =
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   671
let val (Il,r) = Thm.dest_comb imp
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   672
    val (_,imp1) = Thm.dest_comb Il
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   673
    val (Ict1,_) = Thm.dest_comb imp1
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   674
    val (_,ct1) = Thm.dest_comb Ict1
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   675
    val (Ir,_) = Thm.dest_comb r
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   676
    val (_,Ict2r) = Thm.dest_comb Ir
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   677
    val (Ict2,_) = Thm.dest_comb Ict2r
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   678
    val (_,ct2) = Thm.dest_comb Ict2
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   679
in (ct1,ct2) end;
6074
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   680
13186
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   681
fun splitasms asms =
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   682
let fun split(asms',[]) = Tip(rev asms')
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   683
      | split(asms',asm::asms) =
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   684
      let val spl = asm COMP LA_Logic.neqE
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   685
          val (ct1,ct2) = extract(cprop_of spl)
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   686
          val thm1 = assume ct1 and thm2 = assume ct2
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   687
      in Spl(spl,ct1,split(asms',asms@[thm1]),ct2,split(asms',asms@[thm2])) end
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   688
      handle THM _ => split(asm::asms', asms)
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   689
in split([],asms) end;
6074
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   690
13186
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   691
fun fwdproof sg (Tip asms) (j::js) = (mkthm sg asms j, js)
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   692
  | fwdproof sg (Spl(thm,ct1,tree1,ct2,tree2)) js =
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   693
    let val (thm1,js1) = fwdproof sg tree1 js
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   694
        val (thm2,js2) = fwdproof sg tree2 js1
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   695
        val thm1' = implies_intr ct1 thm1
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   696
        val thm2' = implies_intr ct2 thm2
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   697
    in (thm2' COMP (thm1' COMP thm), js2) end;
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   698
(* needs handle _ => None ? *)
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   699
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   700
fun prover sg thms Tconcl js pos =
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   701
let val nTconcl = LA_Logic.neg_prop Tconcl
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   702
    val cnTconcl = cterm_of sg nTconcl
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   703
    val nTconclthm = assume cnTconcl
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   704
    val tree = splitasms (thms @ [nTconclthm])
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   705
    val (thm,_) = fwdproof sg tree js
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   706
    val contr = if pos then LA_Logic.ccontr else LA_Logic.notI
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   707
in Some(LA_Logic.mk_Eq((implies_intr cnTconcl thm) COMP contr)) end
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   708
(* in case concl contains ?-var, which makes assume fail: *)
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   709
handle THM _ => None;
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   710
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   711
(* PRE: concl is not negated!
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   712
   This assumption is OK because
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   713
   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
   714
   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
   715
      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
   716
*)
6074
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   717
fun lin_arith_prover sg thms concl =
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   718
let val Hs = map (#prop o rep_thm) thms
6102
b985e2184868 Split argument structure.
nipkow
parents: 6079
diff changeset
   719
    val Tconcl = LA_Logic.mk_Trueprop concl
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   720
in case prove sg ([],[]) false Hs Tconcl of (* concl provable? *)
13186
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   721
     Some js => prover sg thms Tconcl js true
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   722
   | None => let val nTconcl = LA_Logic.neg_prop Tconcl
13498
5330f1744817 Added counter example generation.
nipkow
parents: 13464
diff changeset
   723
          in case prove sg ([],[]) false Hs nTconcl of (* ~concl provable? *)
13186
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   724
               Some js => prover sg thms nTconcl js false
ef8ed6adcb38 Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents: 13105
diff changeset
   725
             | None => None
6079
4f7975c74cdf More arith simplifications.
nipkow
parents: 6074
diff changeset
   726
          end
5982
aeb97860d352 Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff changeset
   727
end;
6074
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   728
34242451bc91 Added simproc.
nipkow
parents: 6062
diff changeset
   729
end;