author | wenzelm |
Wed, 24 Jul 2002 22:14:42 +0200 | |
changeset 13420 | 39fca1e5818a |
parent 13186 | ef8ed6adcb38 |
child 13464 | c98321b8d638 |
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 |
|
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 neqE: thm (* [| m ~= n; m < n ==> P; n < m ==> P |] ==> P *) |
|
29 |
val notI: thm (* (P ==> False) ==> ~ P *) |
|
6110 | 30 |
val not_lessD: thm (* ~(m < n) ==> n <= m *) |
6128 | 31 |
val not_leD: thm (* ~(m <= n) ==> n < m *) |
6102 | 32 |
val sym: thm (* x = y ==> y = x *) |
33 |
val mk_Eq: thm -> thm |
|
34 |
val mk_Trueprop: term -> term |
|
35 |
val neg_prop: term -> term |
|
36 |
val is_False: thm -> bool |
|
6128 | 37 |
val is_nat: typ list * term -> bool |
38 |
val mk_nat_thm: Sign.sg -> term -> thm |
|
6102 | 39 |
end; |
40 |
(* |
|
41 |
mk_Eq(~in) = `in == False' |
|
42 |
mk_Eq(in) = `in == True' |
|
43 |
where `in' is an (in)equality. |
|
44 |
||
45 |
neg_prop(t) = neg if t is wrapped up in Trueprop and |
|
46 |
nt is the (logically) negated version of t, where the negation |
|
47 |
of a negative term is the term itself (no double negation!); |
|
6128 | 48 |
|
49 |
is_nat(parameter-types,t) = t:nat |
|
50 |
mk_nat_thm(t) = "0 <= t" |
|
6102 | 51 |
*) |
52 |
||
5982
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
53 |
signature LIN_ARITH_DATA = |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
54 |
sig |
6128 | 55 |
val decomp: |
10691 | 56 |
Sign.sg -> term -> ((term*rat)list * rat * string * (term*rat)list * rat * bool)option |
57 |
val number_of: int * typ -> term |
|
5982
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
58 |
end; |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
59 |
(* |
7551
8e934d1a9ac6
Now distinguishes discrete from non-distrete types.
nipkow
parents:
6128
diff
changeset
|
60 |
decomp(`x Rel y') should yield (p,i,Rel,q,j,d) |
5982
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
61 |
where Rel is one of "<", "~<", "<=", "~<=" and "=" and |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
62 |
p/q is the decomposition of the sum terms x/y into a list |
7551
8e934d1a9ac6
Now distinguishes discrete from non-distrete types.
nipkow
parents:
6128
diff
changeset
|
63 |
of summand * multiplicity pairs and a constant summand and |
8e934d1a9ac6
Now distinguishes discrete from non-distrete types.
nipkow
parents:
6128
diff
changeset
|
64 |
d indicates if the domain is discrete. |
5982
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
65 |
|
9420 | 66 |
ss must reduce contradictory <= to False. |
5982
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
67 |
It should also cancel common summands to keep <= reduced; |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
68 |
otherwise <= can grow to massive proportions. |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
69 |
*) |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
70 |
|
6062 | 71 |
signature FAST_LIN_ARITH = |
72 |
sig |
|
9420 | 73 |
val setup: (theory -> theory) list |
10691 | 74 |
val map_data: ({add_mono_thms: thm list, mult_mono_thms: (thm*cterm)list, inj_thms: thm list, |
10575 | 75 |
lessD: thm list, simpset: Simplifier.simpset} |
10691 | 76 |
-> {add_mono_thms: thm list, mult_mono_thms: (thm*cterm)list, inj_thms: thm list, |
10575 | 77 |
lessD: thm list, simpset: Simplifier.simpset}) |
78 |
-> theory -> theory |
|
9073 | 79 |
val trace : bool ref |
6074 | 80 |
val lin_arith_prover: Sign.sg -> thm list -> term -> thm option |
6062 | 81 |
val lin_arith_tac: int -> tactic |
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 |
||
94 |
structure DataArgs = |
|
95 |
struct |
|
96 |
val name = "Provers/fast_lin_arith"; |
|
10691 | 97 |
type T = {add_mono_thms: thm list, mult_mono_thms: (thm*cterm)list, inj_thms: thm list, |
10575 | 98 |
lessD: thm list, simpset: Simplifier.simpset}; |
9420 | 99 |
|
10691 | 100 |
val empty = {add_mono_thms = [], mult_mono_thms = [], inj_thms = [], |
10575 | 101 |
lessD = [], simpset = Simplifier.empty_ss}; |
9420 | 102 |
val copy = I; |
103 |
val prep_ext = I; |
|
104 |
||
10691 | 105 |
fun merge ({add_mono_thms= add_mono_thms1, mult_mono_thms= mult_mono_thms1, inj_thms= inj_thms1, |
10575 | 106 |
lessD = lessD1, simpset = simpset1}, |
10691 | 107 |
{add_mono_thms= add_mono_thms2, mult_mono_thms= mult_mono_thms2, inj_thms= inj_thms2, |
10575 | 108 |
lessD = lessD2, simpset = simpset2}) = |
9420 | 109 |
{add_mono_thms = Drule.merge_rules (add_mono_thms1, add_mono_thms2), |
13105
3d1e7a199bdc
use eq_thm_prop instead of slightly inadequate eq_thm;
wenzelm
parents:
12932
diff
changeset
|
110 |
mult_mono_thms = gen_merge_lists' (Drule.eq_thm_prop o pairself fst) |
3d1e7a199bdc
use eq_thm_prop instead of slightly inadequate eq_thm;
wenzelm
parents:
12932
diff
changeset
|
111 |
mult_mono_thms1 mult_mono_thms2, |
10575 | 112 |
inj_thms = Drule.merge_rules (inj_thms1, inj_thms2), |
113 |
lessD = Drule.merge_rules (lessD1, lessD2), |
|
114 |
simpset = Simplifier.merge_ss (simpset1, simpset2)}; |
|
9420 | 115 |
|
116 |
fun print _ _ = (); |
|
117 |
end; |
|
118 |
||
119 |
structure Data = TheoryDataFun(DataArgs); |
|
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:
6128
diff
changeset
|
141 |
| NotLeDD of injust |
5982
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
142 |
| Multiplied of int * injust |
10691 | 143 |
| Multiplied2 of 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 |
|
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
146 |
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
|
147 |
|
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
148 |
(* ------------------------------------------------------------------------- *) |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
149 |
(* 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
|
150 |
(* ------------------------------------------------------------------------- *) |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
151 |
|
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
152 |
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
|
153 |
| find_add_type(x,Eq) = x |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
154 |
| find_add_type(_,Lt) = Lt |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
155 |
| find_add_type(Lt,_) = Lt |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
156 |
| find_add_type(Le,Le) = Le; |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
157 |
|
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
158 |
(* ------------------------------------------------------------------------- *) |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
159 |
(* Multiply out an (in)equation. *) |
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 |
|
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
162 |
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
|
163 |
if n = 1 then i |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
164 |
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
|
165 |
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
|
166 |
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
|
167 |
|
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 |
(* Add together (in)equations. *) |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
170 |
(* ------------------------------------------------------------------------- *) |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
171 |
|
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
172 |
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
|
173 |
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
|
174 |
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
|
175 |
|
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
176 |
(* ------------------------------------------------------------------------- *) |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
177 |
(* 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
|
178 |
(* 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
|
179 |
(* ------------------------------------------------------------------------- *) |
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 |
fun el 0 (h::_) = h |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
182 |
| 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
|
183 |
| el _ _ = sys_error "el"; |
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 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
|
186 |
let val c1 = el v l1 and c2 = el v l2 |
10691 | 187 |
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
|
188 |
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
|
189 |
val (n1,n2) = |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
190 |
if (c1 >= 0) = (c2 >= 0) |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
191 |
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
|
192 |
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
|
193 |
else sys_error "elim_var" |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
194 |
else (m1,m2) |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
195 |
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
|
196 |
then (~n1,~n2) else (n1,n2) |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
197 |
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
|
198 |
|
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
199 |
(* ------------------------------------------------------------------------- *) |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
200 |
(* The main refutation-finding code. *) |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
201 |
(* ------------------------------------------------------------------------- *) |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
202 |
|
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
203 |
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
|
204 |
|
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
205 |
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
|
206 |
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
|
207 |
|
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
208 |
fun calc_blowup l = |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
209 |
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
|
210 |
in (length p) * (length n) end; |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
211 |
|
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
212 |
(* ------------------------------------------------------------------------- *) |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
213 |
(* Main elimination code: *) |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
214 |
(* *) |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
215 |
(* (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
|
216 |
(* *) |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
217 |
(* (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
|
218 |
(* 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
|
219 |
(* *) |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
220 |
(* (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
|
221 |
(* 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
|
222 |
(* ------------------------------------------------------------------------- *) |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
223 |
|
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
224 |
fun allpairs f xs ys = |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
225 |
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
|
226 |
|
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
227 |
fun extract_first p = |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
228 |
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
|
229 |
else extract (y::xs) ys |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
230 |
| extract xs [] = (None,xs) |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
231 |
in extract [] end; |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
232 |
|
6056 | 233 |
fun print_ineqs ineqs = |
9073 | 234 |
if !trace then |
12262 | 235 |
tracing(cat_lines(""::map (fn Lineq(c,t,l,_) => |
9073 | 236 |
string_of_int c ^ |
237 |
(case t of Eq => " = " | Lt=> " < " | Le => " <= ") ^ |
|
238 |
commas(map string_of_int l)) ineqs)) |
|
239 |
else (); |
|
6056 | 240 |
|
5982
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
241 |
fun elim ineqs = |
9073 | 242 |
let val dummy = print_ineqs ineqs; |
6056 | 243 |
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
|
244 |
if not(null triv) |
13186
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
245 |
then case Library.find_first is_answer triv of |
5982
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
246 |
None => elim nontriv | some => some |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
247 |
else |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
248 |
if null nontriv then None else |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
249 |
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
|
250 |
if not(null eqs) then |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
251 |
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
|
252 |
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
|
253 |
(filter (fn i => i<>0) clist) |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
254 |
val c = hd sclist |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
255 |
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
|
256 |
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
|
257 |
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
|
258 |
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
|
259 |
(othereqs @ noneqs) |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
260 |
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
|
261 |
in elim others end |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
262 |
else |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
263 |
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
|
264 |
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
|
265 |
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
|
266 |
val blows = map calc_blowup coeffs |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
267 |
val iblows = blows ~~ numlist |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
268 |
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
|
269 |
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
|
270 |
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
|
271 |
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
|
272 |
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
|
273 |
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
|
274 |
end |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
275 |
end |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
276 |
end; |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
277 |
|
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
278 |
(* ------------------------------------------------------------------------- *) |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
279 |
(* Translate back a proof. *) |
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 |
|
9073 | 282 |
fun trace_thm msg th = |
12262 | 283 |
if !trace then (tracing msg; tracing (Display.string_of_thm th); th) else th; |
9073 | 284 |
|
285 |
fun trace_msg msg = |
|
12262 | 286 |
if !trace then tracing msg else (); |
9073 | 287 |
|
6056 | 288 |
(* FIXME OPTIMIZE!!!! |
289 |
Addition/Multiplication need i*t representation rather than t+t+... |
|
10691 | 290 |
Get rid of Mulitplied(2). For Nat LA_Data.number_of should return Suc^n |
291 |
because Numerals are not known early enough. |
|
6056 | 292 |
|
293 |
Simplification may detect a contradiction 'prematurely' due to type |
|
294 |
information: n+1 <= 0 is simplified to False and does not need to be crossed |
|
295 |
with 0 <= n. |
|
296 |
*) |
|
297 |
local |
|
298 |
exception FalseE of thm |
|
299 |
in |
|
6074 | 300 |
fun mkthm sg asms just = |
10691 | 301 |
let val {add_mono_thms, mult_mono_thms, inj_thms, lessD, simpset} = Data.get_sg sg; |
9420 | 302 |
val atoms = foldl (fn (ats,(lhs,_,_,rhs,_,_)) => |
6056 | 303 |
map fst lhs union (map fst rhs union ats)) |
9420 | 304 |
([], mapfilter (LA_Data.decomp sg o concl_of) asms) |
6056 | 305 |
|
10575 | 306 |
fun add2 thm1 thm2 = |
6102 | 307 |
let val conj = thm1 RS (thm2 RS LA_Logic.conjI) |
10575 | 308 |
in get_first (fn th => Some(conj RS th) handle _ => None) add_mono_thms |
5982
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
309 |
end; |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
310 |
|
10575 | 311 |
fun try_add [] _ = None |
312 |
| try_add (thm1::thm1s) thm2 = case add2 thm1 thm2 of |
|
313 |
None => try_add thm1s thm2 | some => some; |
|
314 |
||
315 |
fun addthms thm1 thm2 = |
|
316 |
case add2 thm1 thm2 of |
|
317 |
None => (case try_add ([thm1] RL inj_thms) thm2 of |
|
318 |
None => the(try_add ([thm2] RL inj_thms) thm1) |
|
319 |
| Some thm => thm) |
|
320 |
| Some thm => thm; |
|
321 |
||
5982
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
322 |
fun multn(n,thm) = |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
323 |
let fun mul(i,th) = if i=1 then th else mul(i-1, addthms thm th) |
6102 | 324 |
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
|
325 |
|
10691 | 326 |
fun multn2(n,thm) = |
327 |
let val Some(mth,cv) = |
|
328 |
get_first (fn (th,cv) => Some(thm RS th,cv) handle _ => None) mult_mono_thms |
|
329 |
val ct = cterm_of sg (LA_Data.number_of(n,#T(rep_cterm cv))) |
|
330 |
in instantiate ([],[(cv,ct)]) mth end |
|
331 |
||
6056 | 332 |
fun simp thm = |
12932 | 333 |
let val thm' = trace_thm "Simplified:" (full_simplify simpset thm) |
6102 | 334 |
in if LA_Logic.is_False thm' then raise FalseE thm' else thm' end |
6056 | 335 |
|
9073 | 336 |
fun mk(Asm i) = trace_thm "Asm" (nth_elem(i,asms)) |
337 |
| mk(Nat(i)) = (trace_msg "Nat"; |
|
338 |
LA_Logic.mk_nat_thm sg (nth_elem(i,atoms))) |
|
9420 | 339 |
| mk(LessD(j)) = trace_thm "L" (hd([mk j] RL lessD)) |
9073 | 340 |
| mk(NotLeD(j)) = trace_thm "NLe" (mk j RS LA_Logic.not_leD) |
9420 | 341 |
| mk(NotLeDD(j)) = trace_thm "NLeD" (hd([mk j RS LA_Logic.not_leD] RL lessD)) |
9073 | 342 |
| mk(NotLessD(j)) = trace_thm "NL" (mk j RS LA_Logic.not_lessD) |
343 |
| mk(Added(j1,j2)) = simp (trace_thm "+" (addthms (mk j1) (mk j2))) |
|
10717 | 344 |
| mk(Multiplied(n,j)) = (trace_msg("*"^string_of_int n); trace_thm "*" (multn(n,mk j))) |
345 |
| mk(Multiplied2(n,j)) = simp (trace_msg("**"^string_of_int n); trace_thm "**" (multn2(n,mk j))) |
|
5982
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
346 |
|
9073 | 347 |
in trace_msg "mkthm"; |
12932 | 348 |
let val thm = trace_thm "Final thm:" (mk just) |
349 |
in let val fls = simplify simpset thm |
|
13186
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
350 |
in trace_thm "After simplification:" fls; |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
351 |
if LA_Logic.is_False fls then fls |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
352 |
else |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
353 |
(tracing "Assumptions:"; seq print_thm asms; |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
354 |
tracing "Proved:"; print_thm fls; |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
355 |
warning "Linear arithmetic should have refuted the assumptions.\n\ |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
356 |
\Please inform Tobias Nipkow (nipkow@in.tum.de)."; |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
357 |
fls) |
12932 | 358 |
end |
359 |
end handle FalseE thm => (trace_thm "False reached early:" thm; thm) |
|
360 |
end |
|
6056 | 361 |
end; |
5982
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
362 |
|
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
363 |
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
|
364 |
|
10691 | 365 |
fun lcms is = foldl lcm (1,is); |
366 |
||
367 |
fun integ(rlhs,r,rel,rrhs,s,d) = |
|
368 |
let val (rn,rd) = rep_rat r and (sn,sd) = rep_rat s |
|
369 |
val m = lcms(map (abs o snd o rep_rat) (r :: s :: map snd rlhs @ map snd rrhs)) |
|
370 |
fun mult(t,r) = let val (i,j) = rep_rat r in (t,i * (m div j)) end |
|
12932 | 371 |
in (m,(map mult rlhs, rn*(m div rd), rel, map mult rrhs, sn*(m div sd), d)) end |
10691 | 372 |
|
5982
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
373 |
fun mklineq atoms = |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
374 |
let val n = length atoms in |
10691 | 375 |
fn (item,k) => |
376 |
let val (m,(lhs,i,rel,rhs,j,discrete)) = integ item |
|
377 |
val lhsa = map (coeff lhs) atoms |
|
5982
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
378 |
and rhsa = map (coeff rhs) atoms |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
379 |
val diff = map2 (op -) (rhsa,lhsa) |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
380 |
val c = i-j |
6056 | 381 |
val just = Asm k |
13186
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
382 |
fun lineq(c,le,cs,j) = Lineq(c,le,cs, if m=1 then j else Multiplied2(m,j)) |
5982
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
383 |
in case rel of |
10691 | 384 |
"<=" => lineq(c,Le,diff,just) |
7551
8e934d1a9ac6
Now distinguishes discrete from non-distrete types.
nipkow
parents:
6128
diff
changeset
|
385 |
| "~<=" => if discrete |
10691 | 386 |
then lineq(1-c,Le,map (op ~) diff,NotLeDD(just)) |
387 |
else lineq(~c,Lt,map (op ~) diff,NotLeD(just)) |
|
7551
8e934d1a9ac6
Now distinguishes discrete from non-distrete types.
nipkow
parents:
6128
diff
changeset
|
388 |
| "<" => if discrete |
10691 | 389 |
then lineq(c+1,Le,diff,LessD(just)) |
390 |
else lineq(c,Lt,diff,just) |
|
391 |
| "~<" => lineq(~c,Le,map (op~) diff,NotLessD(just)) |
|
392 |
| "=" => lineq(c,Eq,diff,just) |
|
5982
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
393 |
| _ => sys_error("mklineq" ^ rel) |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
394 |
end |
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 |
|
6056 | 397 |
fun mknat pTs ixs (atom,i) = |
6128 | 398 |
if LA_Logic.is_nat(pTs,atom) |
6056 | 399 |
then let val l = map (fn j => if j=i then 1 else 0) ixs |
400 |
in Some(Lineq(0,Le,l,Nat(i))) end |
|
401 |
else None |
|
402 |
||
13186
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
403 |
(* This code is tricky. It takes a list of premises in the order they occur |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
404 |
in the subgoal. Numerical premises are coded as Some(tuple), non-numerical |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
405 |
ones as None. Going through the premises, each numeric one is converted into |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
406 |
a Lineq. The tricky bit is to convert ~= which is split into two cases < and |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
407 |
>. Thus mklineqss returns a list of equation systems. This may blow up if |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
408 |
there are many ~=, but in practice it does not seem to happen. The really |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
409 |
tricky bit is to arrange the order of the cases such that they coincide with |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
410 |
the order in which the cases are in the end generated by the tactic that |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
411 |
applies the generated refutation thms (see function 'refute_tac'). |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
412 |
|
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
413 |
For variables n of type nat, a constraint 0 <= n is added. |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
414 |
*) |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
415 |
fun mklineqss(pTs,items) = |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
416 |
let fun add(ats,None) = ats |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
417 |
| add(ats,Some(lhs,_,_,rhs,_,_)) = |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
418 |
(map fst lhs) union ((map fst rhs) union ats) |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
419 |
val atoms = foldl add ([],items) |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
420 |
val mkleq = mklineq atoms |
6056 | 421 |
val ixs = 0 upto (length(atoms)-1) |
422 |
val iatoms = atoms ~~ ixs |
|
13186
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
423 |
val natlineqs = mapfilter (mknat pTs ixs) iatoms |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
424 |
|
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
425 |
fun elim_neq front _ [] = [front] |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
426 |
| elim_neq front n (None::ineqs) = elim_neq front (n+1) ineqs |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
427 |
| elim_neq front n (Some(ineq as (l,i,rel,r,j,d))::ineqs) = |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
428 |
if rel = "~=" then elim_neq front n (ineqs @ [Some(l,i,"<",r,j,d)]) @ |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
429 |
elim_neq front n (ineqs @ [Some(r,j,"<",l,i,d)]) |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
430 |
else elim_neq (mkleq(ineq,n) :: front) (n+1) ineqs |
5982
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
431 |
|
13186
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
432 |
in elim_neq natlineqs 0 items end; |
6074 | 433 |
|
13186
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
434 |
fun elim_all (ineqs::ineqss) js = |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
435 |
(case elim ineqs of None => None |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
436 |
| Some(Lineq(_,_,_,j)) => elim_all ineqss (js@[j])) |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
437 |
| elim_all [] js = Some js |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
438 |
|
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
439 |
fun refute(pTsitems) = elim_all (mklineqss pTsitems) []; |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
440 |
|
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
441 |
fun refute_tac(i,justs) = |
6074 | 442 |
fn state => |
443 |
let val sg = #sign(rep_thm state) |
|
13186
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
444 |
fun just1 j = REPEAT_DETERM(etac LA_Logic.neqE i) THEN |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
445 |
METAHYPS (fn asms => rtac (mkthm sg asms j) 1) i |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
446 |
in DETERM(resolve_tac [LA_Logic.notI,LA_Logic.ccontr] i) THEN |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
447 |
EVERY(map just1 justs) |
6074 | 448 |
end |
449 |
state; |
|
450 |
||
9420 | 451 |
fun prove sg (pTs,Hs,concl) = |
13186
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
452 |
let val Hitems = map (LA_Data.decomp sg) Hs |
9420 | 453 |
in case LA_Data.decomp sg concl of |
13186
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
454 |
None => refute(pTs,Hitems@[None]) |
7551
8e934d1a9ac6
Now distinguishes discrete from non-distrete types.
nipkow
parents:
6128
diff
changeset
|
455 |
| Some(citem as (r,i,rel,l,j,d)) => |
13186
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
456 |
let val neg::rel0 = explode rel |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
457 |
val nrel = if neg = "~" then implode rel0 else "~"^rel |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
458 |
in refute(pTs, Hitems @ [Some(r,i,nrel,l,j,d)]) end |
6074 | 459 |
end; |
5982
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
460 |
|
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
461 |
(* |
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
462 |
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
|
463 |
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
|
464 |
*) |
13186
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
465 |
fun lin_arith_tac i st = SUBGOAL (fn (A,_) => |
6056 | 466 |
let val pTs = rev(map snd (Logic.strip_params A)) |
467 |
val Hs = Logic.strip_assums_hyp A |
|
6074 | 468 |
val concl = Logic.strip_assums_concl A |
12932 | 469 |
in trace_thm ("Trying to refute subgoal " ^ string_of_int i) st; |
470 |
case prove (Thm.sign_of_thm st) (pTs,Hs,concl) of |
|
13186
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
471 |
None => no_tac |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
472 |
| Some js => refute_tac(i,js) |
9420 | 473 |
end) i st; |
5982
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
474 |
|
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
475 |
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
|
476 |
|
13186
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
477 |
(** Forward proof from theorems **) |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
478 |
|
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
479 |
(* More tricky code. Needs to arrange the proofs of the multiple cases (due |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
480 |
to splits of ~= premises) such that it coincides with the order of the cases |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
481 |
generated by function mklineqss. *) |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
482 |
|
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
483 |
datatype splittree = Tip of thm list |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
484 |
| Spl of thm * cterm * splittree * cterm * splittree |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
485 |
|
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
486 |
fun extract imp = |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
487 |
let val (Il,r) = Thm.dest_comb imp |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
488 |
val (_,imp1) = Thm.dest_comb Il |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
489 |
val (Ict1,_) = Thm.dest_comb imp1 |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
490 |
val (_,ct1) = Thm.dest_comb Ict1 |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
491 |
val (Ir,_) = Thm.dest_comb r |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
492 |
val (_,Ict2r) = Thm.dest_comb Ir |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
493 |
val (Ict2,_) = Thm.dest_comb Ict2r |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
494 |
val (_,ct2) = Thm.dest_comb Ict2 |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
495 |
in (ct1,ct2) end; |
6074 | 496 |
|
13186
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
497 |
fun splitasms asms = |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
498 |
let fun split(asms',[]) = Tip(rev asms') |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
499 |
| split(asms',asm::asms) = |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
500 |
let val spl = asm COMP LA_Logic.neqE |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
501 |
val (ct1,ct2) = extract(cprop_of spl) |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
502 |
val thm1 = assume ct1 and thm2 = assume ct2 |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
503 |
in Spl(spl,ct1,split(asms',asms@[thm1]),ct2,split(asms',asms@[thm2])) end |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
504 |
handle THM _ => split(asm::asms', asms) |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
505 |
in split([],asms) end; |
6074 | 506 |
|
13186
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
507 |
fun fwdproof sg (Tip asms) (j::js) = (mkthm sg asms j, js) |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
508 |
| fwdproof sg (Spl(thm,ct1,tree1,ct2,tree2)) js = |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
509 |
let val (thm1,js1) = fwdproof sg tree1 js |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
510 |
val (thm2,js2) = fwdproof sg tree2 js1 |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
511 |
val thm1' = implies_intr ct1 thm1 |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
512 |
val thm2' = implies_intr ct2 thm2 |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
513 |
in (thm2' COMP (thm1' COMP thm), js2) end; |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
514 |
(* needs handle _ => None ? *) |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
515 |
|
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
516 |
fun prover sg thms Tconcl js pos = |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
517 |
let val nTconcl = LA_Logic.neg_prop Tconcl |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
518 |
val cnTconcl = cterm_of sg nTconcl |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
519 |
val nTconclthm = assume cnTconcl |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
520 |
val tree = splitasms (thms @ [nTconclthm]) |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
521 |
val (thm,_) = fwdproof sg tree js |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
522 |
val contr = if pos then LA_Logic.ccontr else LA_Logic.notI |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
523 |
in Some(LA_Logic.mk_Eq((implies_intr cnTconcl thm) COMP contr)) end |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
524 |
(* in case concl contains ?-var, which makes assume fail: *) |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
525 |
handle THM _ => None; |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
526 |
|
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
527 |
(* PRE: concl is not negated! |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
528 |
This assumption is OK because |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
529 |
1. lin_arith_prover tries both to prove and disprove concl and |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
530 |
2. lin_arith_prover is applied by the simplifier which |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
531 |
dives into terms and will thus try the non-negated concl anyway. |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
532 |
*) |
6074 | 533 |
fun lin_arith_prover sg thms concl = |
534 |
let val Hs = map (#prop o rep_thm) thms |
|
6102 | 535 |
val Tconcl = LA_Logic.mk_Trueprop concl |
13186
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
536 |
in case prove sg ([],Hs,Tconcl) of (* concl provable? *) |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
537 |
Some js => prover sg thms Tconcl js true |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
538 |
| None => let val nTconcl = LA_Logic.neg_prop Tconcl |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
539 |
in case prove sg ([],Hs,nTconcl) of (* ~concl provable? *) |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
540 |
Some js => prover sg thms nTconcl js false |
ef8ed6adcb38
Big update. Allows case splitting on ~= now (trying both < and >).
nipkow
parents:
13105
diff
changeset
|
541 |
| None => None |
6079 | 542 |
end |
5982
aeb97860d352
Replaced the puny nat_transitive.ML by the general fast_lin_arith.ML.
nipkow
parents:
diff
changeset
|
543 |
end; |
6074 | 544 |
|
545 |
end; |