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