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