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