author | wenzelm |
Wed, 14 Oct 2015 18:29:41 +0200 | |
changeset 61442 | 467ebb937294 |
parent 61144 | 5e94dfead1c2 |
child 61841 | 4d3527b94f2a |
permissions | -rw-r--r-- |
37744 | 1 |
(* Title: HOL/Tools/groebner.ML |
23252 | 2 |
Author: Amine Chaieb, TU Muenchen |
3 |
*) |
|
4 |
||
5 |
signature GROEBNER = |
|
6 |
sig |
|
41453 | 7 |
val ring_and_ideal_conv: |
30866 | 8 |
{idom: thm list, ring: cterm list * thm list, field: cterm list * thm list, |
9 |
vars: cterm list, semiring: cterm list * thm list, ideal : thm list} -> |
|
23487 | 10 |
(cterm -> Rat.rat) -> (Rat.rat -> cterm) -> |
11 |
conv -> conv -> |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47432
diff
changeset
|
12 |
{ring_conv: Proof.context -> conv, |
36723 | 13 |
simple_ideal: (cterm list -> cterm -> (cterm * cterm -> order) -> cterm list), |
14 |
multi_ideal: cterm list -> cterm list -> cterm list -> (cterm * cterm) list, |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47432
diff
changeset
|
15 |
poly_eq_ss: simpset, unwind_conv: Proof.context -> conv} |
36723 | 16 |
val ring_tac: thm list -> thm list -> Proof.context -> int -> tactic |
17 |
val ideal_tac: thm list -> thm list -> Proof.context -> int -> tactic |
|
18 |
val algebra_tac: thm list -> thm list -> Proof.context -> int -> tactic |
|
23252 | 19 |
end |
20 |
||
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
21 |
structure Groebner : GROEBNER = |
23252 | 22 |
struct |
23 |
||
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
24 |
val concl = Thm.cprop_of #> Thm.dest_arg; |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
25 |
|
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
26 |
fun is_binop ct ct' = |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
27 |
(case Thm.term_of ct' of |
59582 | 28 |
c $ _ $ _ => Thm.term_of ct aconv c |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
29 |
| _ => false); |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
30 |
|
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
31 |
fun dest_binary ct ct' = |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
32 |
if is_binop ct ct' then Thm.dest_binop ct' |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
33 |
else raise CTERM ("dest_binary: bad binop", [ct, ct']) |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
34 |
|
23252 | 35 |
val rat_0 = Rat.zero; |
36 |
val rat_1 = Rat.one; |
|
37 |
val minus_rat = Rat.neg; |
|
38 |
val denominator_rat = Rat.quotient_of_rat #> snd #> Rat.rat_of_int; |
|
39 |
fun int_of_rat a = |
|
40 |
case Rat.quotient_of_rat a of (i,1) => i | _ => error "int_of_rat: not an int"; |
|
23514 | 41 |
val lcm_rat = fn x => fn y => Rat.rat_of_int (Integer.lcm (int_of_rat x) (int_of_rat y)); |
23252 | 42 |
|
43 |
val (eqF_intr, eqF_elim) = |
|
36713 | 44 |
let val [th1,th2] = @{thms PFalse} |
23252 | 45 |
in (fn th => th COMP th2, fn th => th COMP th1) end; |
46 |
||
47 |
val (PFalse, PFalse') = |
|
36713 | 48 |
let val PFalse_eq = nth @{thms simp_thms} 13 |
23252 | 49 |
in (PFalse_eq RS iffD1, PFalse_eq RS iffD2) end; |
50 |
||
51 |
||
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
52 |
(* Type for recording history, i.e. how a polynomial was obtained. *) |
23252 | 53 |
|
54 |
datatype history = |
|
24630
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
wenzelm
parents:
23579
diff
changeset
|
55 |
Start of int |
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
wenzelm
parents:
23579
diff
changeset
|
56 |
| Mmul of (Rat.rat * int list) * history |
23252 | 57 |
| Add of history * history; |
58 |
||
59 |
||
60 |
(* Monomial ordering. *) |
|
61 |
||
62 |
fun morder_lt m1 m2= |
|
63 |
let fun lexorder l1 l2 = |
|
64 |
case (l1,l2) of |
|
65 |
([],[]) => false |
|
66 |
| (x1::o1,x2::o2) => x1 > x2 orelse x1 = x2 andalso lexorder o1 o2 |
|
67 |
| _ => error "morder: inconsistent monomial lengths" |
|
24630
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
wenzelm
parents:
23579
diff
changeset
|
68 |
val n1 = Integer.sum m1 |
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
wenzelm
parents:
23579
diff
changeset
|
69 |
val n2 = Integer.sum m2 in |
23252 | 70 |
n1 < n2 orelse n1 = n2 andalso lexorder m1 m2 |
71 |
end; |
|
72 |
||
73 |
(* Arithmetic on canonical polynomials. *) |
|
74 |
||
75 |
fun grob_neg l = map (fn (c,m) => (minus_rat c,m)) l; |
|
76 |
||
77 |
fun grob_add l1 l2 = |
|
78 |
case (l1,l2) of |
|
79 |
([],l2) => l2 |
|
80 |
| (l1,[]) => l1 |
|
81 |
| ((c1,m1)::o1,(c2,m2)::o2) => |
|
82 |
if m1 = m2 then |
|
83 |
let val c = c1+/c2 val rest = grob_add o1 o2 in |
|
84 |
if c =/ rat_0 then rest else (c,m1)::rest end |
|
85 |
else if morder_lt m2 m1 then (c1,m1)::(grob_add o1 l2) |
|
86 |
else (c2,m2)::(grob_add l1 o2); |
|
87 |
||
88 |
fun grob_sub l1 l2 = grob_add l1 (grob_neg l2); |
|
89 |
||
24630
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
wenzelm
parents:
23579
diff
changeset
|
90 |
fun grob_mmul (c1,m1) (c2,m2) = (c1*/c2, ListPair.map (op +) (m1, m2)); |
23252 | 91 |
|
92 |
fun grob_cmul cm pol = map (grob_mmul cm) pol; |
|
93 |
||
94 |
fun grob_mul l1 l2 = |
|
95 |
case l1 of |
|
96 |
[] => [] |
|
97 |
| (h1::t1) => grob_add (grob_cmul h1 l2) (grob_mul t1 l2); |
|
98 |
||
99 |
fun grob_inv l = |
|
100 |
case l of |
|
101 |
[(c,vs)] => if (forall (fn x => x = 0) vs) then |
|
102 |
if (c =/ rat_0) then error "grob_inv: division by zero" |
|
103 |
else [(rat_1 // c,vs)] |
|
104 |
else error "grob_inv: non-constant divisor polynomial" |
|
105 |
| _ => error "grob_inv: non-constant divisor polynomial"; |
|
106 |
||
107 |
fun grob_div l1 l2 = |
|
108 |
case l2 of |
|
109 |
[(c,l)] => if (forall (fn x => x = 0) l) then |
|
110 |
if c =/ rat_0 then error "grob_div: division by zero" |
|
111 |
else grob_cmul (rat_1 // c,l) l1 |
|
112 |
else error "grob_div: non-constant divisor polynomial" |
|
113 |
| _ => error "grob_div: non-constant divisor polynomial"; |
|
114 |
||
115 |
fun grob_pow vars l n = |
|
116 |
if n < 0 then error "grob_pow: negative power" |
|
54251 | 117 |
else if n = 0 then [(rat_1,map (K 0) vars)] |
24630
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
wenzelm
parents:
23579
diff
changeset
|
118 |
else grob_mul l (grob_pow vars l (n - 1)); |
23252 | 119 |
|
120 |
(* Monomial division operation. *) |
|
121 |
||
122 |
fun mdiv (c1,m1) (c2,m2) = |
|
123 |
(c1//c2, |
|
24630
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
wenzelm
parents:
23579
diff
changeset
|
124 |
map2 (fn n1 => fn n2 => if n1 < n2 then error "mdiv" else n1 - n2) m1 m2); |
23252 | 125 |
|
126 |
(* Lowest common multiple of two monomials. *) |
|
127 |
||
54251 | 128 |
fun mlcm (_,m1) (_,m2) = (rat_1, ListPair.map Int.max (m1, m2)); |
23252 | 129 |
|
130 |
(* Reduce monomial cm by polynomial pol, returning replacement for cm. *) |
|
131 |
||
132 |
fun reduce1 cm (pol,hpol) = |
|
133 |
case pol of |
|
134 |
[] => error "reduce1" |
|
135 |
| cm1::cms => ((let val (c,m) = mdiv cm cm1 in |
|
136 |
(grob_cmul (minus_rat c,m) cms, |
|
137 |
Mmul((minus_rat c,m),hpol)) end) |
|
138 |
handle ERROR _ => error "reduce1"); |
|
139 |
||
140 |
(* Try this for all polynomials in a basis. *) |
|
141 |
fun tryfind f l = |
|
142 |
case l of |
|
143 |
[] => error "tryfind" |
|
144 |
| (h::t) => ((f h) handle ERROR _ => tryfind f t); |
|
145 |
||
146 |
fun reduceb cm basis = tryfind (fn p => reduce1 cm p) basis; |
|
147 |
||
148 |
(* Reduction of a polynomial (always picking largest monomial possible). *) |
|
149 |
||
150 |
fun reduce basis (pol,hist) = |
|
151 |
case pol of |
|
152 |
[] => (pol,hist) |
|
153 |
| cm::ptl => ((let val (q,hnew) = reduceb cm basis in |
|
154 |
reduce basis (grob_add q ptl,Add(hnew,hist)) end) |
|
155 |
handle (ERROR _) => |
|
156 |
(let val (q,hist') = reduce basis (ptl,hist) in |
|
157 |
(cm::q,hist') end)); |
|
158 |
||
159 |
(* Check for orthogonality w.r.t. LCM. *) |
|
160 |
||
161 |
fun orthogonal l p1 p2 = |
|
162 |
snd l = snd(grob_mmul (hd p1) (hd p2)); |
|
163 |
||
164 |
(* Compute S-polynomial of two polynomials. *) |
|
165 |
||
166 |
fun spoly cm ph1 ph2 = |
|
167 |
case (ph1,ph2) of |
|
54251 | 168 |
(([],h),_) => ([],h) |
169 |
| (_,([],h)) => ([],h) |
|
23252 | 170 |
| ((cm1::ptl1,his1),(cm2::ptl2,his2)) => |
171 |
(grob_sub (grob_cmul (mdiv cm cm1) ptl1) |
|
172 |
(grob_cmul (mdiv cm cm2) ptl2), |
|
173 |
Add(Mmul(mdiv cm cm1,his1), |
|
174 |
Mmul(mdiv (minus_rat(fst cm),snd cm) cm2,his2))); |
|
175 |
||
176 |
(* Make a polynomial monic. *) |
|
177 |
||
178 |
fun monic (pol,hist) = |
|
23579 | 179 |
if null pol then (pol,hist) else |
23252 | 180 |
let val (c',m') = hd pol in |
181 |
(map (fn (c,m) => (c//c',m)) pol, |
|
182 |
Mmul((rat_1 // c',map (K 0) m'),hist)) end; |
|
183 |
||
184 |
(* The most popular heuristic is to order critical pairs by LCM monomial. *) |
|
185 |
||
54251 | 186 |
fun forder ((_,m1),_) ((_,m2),_) = morder_lt m1 m2; |
23252 | 187 |
|
188 |
fun poly_lt p q = |
|
189 |
case (p,q) of |
|
54251 | 190 |
(_,[]) => false |
191 |
| ([],_) => true |
|
23252 | 192 |
| ((c1,m1)::o1,(c2,m2)::o2) => |
193 |
c1 </ c2 orelse |
|
194 |
c1 =/ c2 andalso ((morder_lt m1 m2) orelse m1 = m2 andalso poly_lt o1 o2); |
|
195 |
||
196 |
fun align ((p,hp),(q,hq)) = |
|
197 |
if poly_lt p q then ((p,hp),(q,hq)) else ((q,hq),(p,hp)); |
|
198 |
||
199 |
fun poly_eq p1 p2 = |
|
40718 | 200 |
eq_list (fn ((c1, m1), (c2, m2)) => c1 =/ c2 andalso (m1: int list) = m2) (p1, p2); |
23252 | 201 |
|
54251 | 202 |
fun memx ((p1,_),(p2,_)) ppairs = |
23252 | 203 |
not (exists (fn ((q1,_),(q2,_)) => poly_eq p1 q1 andalso poly_eq p2 q2) ppairs); |
204 |
||
205 |
(* Buchberger's second criterion. *) |
|
206 |
||
207 |
fun criterion2 basis (lcm,((p1,h1),(p2,h2))) opairs = |
|
208 |
exists (fn g => not(poly_eq (fst g) p1) andalso not(poly_eq (fst g) p2) andalso |
|
209 |
can (mdiv lcm) (hd(fst g)) andalso |
|
210 |
not(memx (align (g,(p1,h1))) (map snd opairs)) andalso |
|
211 |
not(memx (align (g,(p2,h2))) (map snd opairs))) basis; |
|
212 |
||
213 |
(* Test for hitting constant polynomial. *) |
|
214 |
||
215 |
fun constant_poly p = |
|
24630
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
wenzelm
parents:
23579
diff
changeset
|
216 |
length p = 1 andalso forall (fn x => x = 0) (snd(hd p)); |
23252 | 217 |
|
218 |
(* Grobner basis algorithm. *) |
|
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
219 |
|
23252 | 220 |
(* FIXME: try to get rid of mergesort? *) |
221 |
fun merge ord l1 l2 = |
|
222 |
case l1 of |
|
223 |
[] => l2 |
|
224 |
| h1::t1 => |
|
225 |
case l2 of |
|
226 |
[] => l1 |
|
227 |
| h2::t2 => if ord h1 h2 then h1::(merge ord t1 l2) |
|
228 |
else h2::(merge ord l1 t2); |
|
229 |
fun mergesort ord l = |
|
230 |
let |
|
231 |
fun mergepairs l1 l2 = |
|
232 |
case (l1,l2) of |
|
233 |
([s],[]) => s |
|
234 |
| (l,[]) => mergepairs [] l |
|
235 |
| (l,[s1]) => mergepairs (s1::l) [] |
|
236 |
| (l,(s1::s2::ss)) => mergepairs ((merge ord s1 s2)::l) ss |
|
23579 | 237 |
in if null l then [] else mergepairs [] (map (fn x => [x]) l) |
23252 | 238 |
end; |
239 |
||
240 |
||
241 |
fun grobner_basis basis pairs = |
|
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
242 |
case pairs of |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
243 |
[] => basis |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
244 |
| (l,(p1,p2))::opairs => |
54251 | 245 |
let val (sph as (sp,_)) = monic (reduce basis (spoly l p1 p2)) |
41453 | 246 |
in |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
247 |
if null sp orelse criterion2 basis (l,(p1,p2)) opairs |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
248 |
then grobner_basis basis opairs |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
249 |
else if constant_poly sp then grobner_basis (sph::basis) [] |
41453 | 250 |
else |
251 |
let |
|
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
252 |
val rawcps = map (fn p => (mlcm (hd(fst p)) (hd sp),align(p,sph))) |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
253 |
basis |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
254 |
val newcps = filter (fn (l,(p,q)) => not(orthogonal l (fst p) (fst q))) |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
255 |
rawcps |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
256 |
in grobner_basis (sph::basis) |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
257 |
(merge forder opairs (mergesort forder newcps)) |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
258 |
end |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
259 |
end; |
23252 | 260 |
|
261 |
(* Interreduce initial polynomials. *) |
|
262 |
||
263 |
fun grobner_interreduce rpols ipols = |
|
264 |
case ipols of |
|
265 |
[] => map monic (rev rpols) |
|
266 |
| p::ps => let val p' = reduce (rpols @ ps) p in |
|
23579 | 267 |
if null (fst p') then grobner_interreduce rpols ps |
23252 | 268 |
else grobner_interreduce (p'::rpols) ps end; |
269 |
||
270 |
(* Overall function. *) |
|
271 |
||
272 |
fun grobner pols = |
|
33063 | 273 |
let val npols = map_index (fn (n, p) => (p, Start n)) pols |
23579 | 274 |
val phists = filter (fn (p,_) => not (null p)) npols |
23252 | 275 |
val bas = grobner_interreduce [] (map monic phists) |
25538 | 276 |
val prs0 = map_product pair bas bas |
23252 | 277 |
val prs1 = filter (fn ((x,_),(y,_)) => poly_lt x y) prs0 |
278 |
val prs2 = map (fn (p,q) => (mlcm (hd(fst p)) (hd(fst q)),(p,q))) prs1 |
|
279 |
val prs3 = |
|
280 |
filter (fn (l,(p,q)) => not(orthogonal l (fst p) (fst q))) prs2 in |
|
281 |
grobner_basis bas (mergesort forder prs3) end; |
|
282 |
||
283 |
(* Get proof of contradiction from Grobner basis. *) |
|
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
284 |
|
23252 | 285 |
fun find p l = |
286 |
case l of |
|
287 |
[] => error "find" |
|
288 |
| (h::t) => if p(h) then h else find p t; |
|
289 |
||
290 |
fun grobner_refute pols = |
|
291 |
let val gb = grobner pols in |
|
54251 | 292 |
snd(find (fn (p,_) => length p = 1 andalso forall (fn x=> x=0) (snd(hd p))) gb) |
23252 | 293 |
end; |
294 |
||
295 |
(* Turn proof into a certificate as sum of multipliers. *) |
|
296 |
(* In principle this is very inefficient: in a heavily shared proof it may *) |
|
297 |
(* make the same calculation many times. Could put in a cache or something. *) |
|
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
298 |
|
23252 | 299 |
fun resolve_proof vars prf = |
300 |
case prf of |
|
301 |
Start(~1) => [] |
|
302 |
| Start m => [(m,[(rat_1,map (K 0) vars)])] |
|
303 |
| Mmul(pol,lin) => |
|
304 |
let val lis = resolve_proof vars lin in |
|
305 |
map (fn (n,p) => (n,grob_cmul pol p)) lis end |
|
306 |
| Add(lin1,lin2) => |
|
307 |
let val lis1 = resolve_proof vars lin1 |
|
308 |
val lis2 = resolve_proof vars lin2 |
|
33042 | 309 |
val dom = distinct (op =) (union (op =) (map fst lis1) (map fst lis2)) |
23252 | 310 |
in |
23557 | 311 |
map (fn n => let val a = these (AList.lookup (op =) lis1 n) |
312 |
val b = these (AList.lookup (op =) lis2 n) |
|
313 |
in (n,grob_add a b) end) dom end; |
|
23252 | 314 |
|
315 |
(* Run the procedure and produce Weak Nullstellensatz certificate. *) |
|
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
316 |
|
23252 | 317 |
fun grobner_weak vars pols = |
318 |
let val cert = resolve_proof vars (grobner_refute pols) |
|
319 |
val l = |
|
320 |
fold_rev (fold_rev (lcm_rat o denominator_rat o fst) o snd) cert (rat_1) in |
|
321 |
(l,map (fn (i,p) => (i,map (fn (d,m) => (l*/d,m)) p)) cert) end; |
|
322 |
||
323 |
(* Prove a polynomial is in ideal generated by others, using Grobner basis. *) |
|
324 |
||
325 |
fun grobner_ideal vars pols pol = |
|
326 |
let val (pol',h) = reduce (grobner pols) (grob_neg pol,Start(~1)) in |
|
24913 | 327 |
if not (null pol') then error "grobner_ideal: not in the ideal" else |
23252 | 328 |
resolve_proof vars h end; |
329 |
||
330 |
(* Produce Strong Nullstellensatz certificate for a power of pol. *) |
|
331 |
||
332 |
fun grobner_strong vars pols pol = |
|
333 |
let val vars' = @{cterm "True"}::vars |
|
54251 | 334 |
val grob_z = [(rat_1,1::(map (K 0) vars))] |
335 |
val grob_1 = [(rat_1,(map (K 0) vars'))] |
|
23252 | 336 |
fun augment p= map (fn (c,m) => (c,0::m)) p |
337 |
val pols' = map augment pols |
|
338 |
val pol' = augment pol |
|
339 |
val allpols = (grob_sub (grob_mul grob_z pol') grob_1)::pols' |
|
340 |
val (l,cert) = grobner_weak vars' allpols |
|
33029 | 341 |
val d = fold (fold (Integer.max o hd o snd) o snd) cert 0 |
23252 | 342 |
fun transform_monomial (c,m) = |
343 |
grob_cmul (c,tl m) (grob_pow vars pol (d - hd m)) |
|
344 |
fun transform_polynomial q = fold_rev (grob_add o transform_monomial) q [] |
|
345 |
val cert' = map (fn (c,q) => (c-1,transform_polynomial q)) |
|
346 |
(filter (fn (k,_) => k <> 0) cert) in |
|
347 |
(d,l,cert') end; |
|
348 |
||
349 |
||
350 |
(* Overall parametrized universal procedure for (semi)rings. *) |
|
351 |
(* We return an ideal_conv and the actual ring prover. *) |
|
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
352 |
|
23252 | 353 |
fun refute_disj rfn tm = |
59582 | 354 |
case Thm.term_of tm of |
54251 | 355 |
Const(@{const_name HOL.disj},_)$_$_ => |
52467 | 356 |
Drule.compose |
41453 | 357 |
(refute_disj rfn (Thm.dest_arg tm), 2, |
52467 | 358 |
Drule.compose (refute_disj rfn (Thm.dest_arg1 tm), 2, disjE)) |
23252 | 359 |
| _ => rfn tm ; |
360 |
||
36713 | 361 |
val notnotD = @{thm notnotD}; |
46497
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
wenzelm
parents:
42793
diff
changeset
|
362 |
fun mk_binop ct x y = Thm.apply (Thm.apply ct x) y |
23252 | 363 |
|
364 |
fun is_neg t = |
|
59582 | 365 |
case Thm.term_of t of |
54251 | 366 |
(Const(@{const_name Not},_)$_) => true |
23252 | 367 |
| _ => false; |
368 |
fun is_eq t = |
|
59582 | 369 |
case Thm.term_of t of |
38864
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
haftmann
parents:
38795
diff
changeset
|
370 |
(Const(@{const_name HOL.eq},_)$_$_) => true |
23252 | 371 |
| _ => false; |
372 |
||
373 |
fun end_itlist f l = |
|
374 |
case l of |
|
375 |
[] => error "end_itlist" |
|
376 |
| [x] => x |
|
377 |
| (h::t) => f h (end_itlist f t); |
|
378 |
||
379 |
val list_mk_binop = fn b => end_itlist (mk_binop b); |
|
380 |
||
381 |
val list_dest_binop = fn b => |
|
382 |
let fun h acc t = |
|
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
383 |
((let val (l,r) = dest_binary b t in h (h acc r) l end) |
23252 | 384 |
handle CTERM _ => (t::acc)) (* Why had I handle _ => ? *) |
385 |
in h [] |
|
386 |
end; |
|
387 |
||
388 |
val strip_exists = |
|
389 |
let fun h (acc, t) = |
|
59582 | 390 |
case Thm.term_of t of |
54251 | 391 |
Const (@{const_name Ex}, _) $ Abs _ => |
41453 | 392 |
h (Thm.dest_abs NONE (Thm.dest_arg t) |>> (fn v => v::acc)) |
23252 | 393 |
| _ => (acc,t) |
394 |
in fn t => h ([],t) |
|
395 |
end; |
|
396 |
||
397 |
fun is_forall t = |
|
59582 | 398 |
case Thm.term_of t of |
38558 | 399 |
(Const(@{const_name All},_)$Abs(_,_,_)) => true |
23252 | 400 |
| _ => false; |
401 |
||
402 |
val mk_object_eq = fn th => th COMP meta_eq_to_obj_eq; |
|
36713 | 403 |
val nnf_simps = @{thms nnf_simps}; |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47432
diff
changeset
|
404 |
|
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47432
diff
changeset
|
405 |
fun weak_dnf_conv ctxt = |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47432
diff
changeset
|
406 |
Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms weak_dnf_simps}); |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47432
diff
changeset
|
407 |
|
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47432
diff
changeset
|
408 |
val initial_ss = |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47432
diff
changeset
|
409 |
simpset_of (put_simpset HOL_basic_ss @{context} |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47432
diff
changeset
|
410 |
addsimps nnf_simps |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47432
diff
changeset
|
411 |
addsimps [not_all, not_ex] |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47432
diff
changeset
|
412 |
addsimps map (fn th => th RS sym) (@{thms ex_simps} @ @{thms all_simps})); |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47432
diff
changeset
|
413 |
fun initial_conv ctxt = |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47432
diff
changeset
|
414 |
Simplifier.rewrite (put_simpset initial_ss ctxt); |
23252 | 415 |
|
60801 | 416 |
val specl = fold_rev (fn x => fn th => Thm.instantiate' [] [SOME x] (th RS spec)); |
23252 | 417 |
|
418 |
val cTrp = @{cterm "Trueprop"}; |
|
38795
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents:
38786
diff
changeset
|
419 |
val cConj = @{cterm HOL.conj}; |
23252 | 420 |
val (cNot,false_tm) = (@{cterm "Not"}, @{cterm "False"}); |
46497
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
wenzelm
parents:
42793
diff
changeset
|
421 |
val assume_Trueprop = Thm.apply cTrp #> Thm.assume; |
23252 | 422 |
val list_mk_conj = list_mk_binop cConj; |
423 |
val conjs = list_dest_binop cConj; |
|
46497
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
wenzelm
parents:
42793
diff
changeset
|
424 |
val mk_neg = Thm.apply cNot; |
23252 | 425 |
|
41453 | 426 |
fun striplist dest = |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
427 |
let |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
428 |
fun h acc x = case try dest x of |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
429 |
SOME (a,b) => h (h acc b) a |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
430 |
| NONE => x::acc |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
431 |
in h [] end; |
46497
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
wenzelm
parents:
42793
diff
changeset
|
432 |
fun list_mk_binop b = foldr1 (fn (s,t) => Thm.apply (Thm.apply b s) t); |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
433 |
|
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
434 |
val eq_commute = mk_meta_eq @{thm eq_commute}; |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
435 |
|
41453 | 436 |
fun sym_conv eq = |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
437 |
let val (l,r) = Thm.dest_binop eq |
60801 | 438 |
in Thm.instantiate' [SOME (Thm.ctyp_of_cterm l)] [SOME l, SOME r] eq_commute |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
439 |
end; |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
440 |
|
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
441 |
(* FIXME : copied from cqe.ML -- complex QE*) |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
442 |
fun conjuncts ct = |
59582 | 443 |
case Thm.term_of ct of |
38795
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents:
38786
diff
changeset
|
444 |
@{term HOL.conj}$_$_ => (Thm.dest_arg1 ct)::(conjuncts (Thm.dest_arg ct)) |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
445 |
| _ => [ct]; |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
446 |
|
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
447 |
fun fold1 f = foldr1 (uncurry f); |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
448 |
|
41453 | 449 |
fun mk_conj_tab th = |
450 |
let fun h acc th = |
|
59582 | 451 |
case Thm.prop_of th of |
54251 | 452 |
@{term "Trueprop"}$(@{term HOL.conj}$_$_) => |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
453 |
h (h acc (th RS conjunct2)) (th RS conjunct1) |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
454 |
| @{term "Trueprop"}$p => (p,th)::acc |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
455 |
in fold (Termtab.insert Thm.eq_thm) (h [] th) Termtab.empty end; |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
456 |
|
38795
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents:
38786
diff
changeset
|
457 |
fun is_conj (@{term HOL.conj}$_$_) = true |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
458 |
| is_conj _ = false; |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
459 |
|
41453 | 460 |
fun prove_conj tab cjs = |
461 |
case cjs of |
|
59582 | 462 |
[c] => if is_conj (Thm.term_of c) then prove_conj tab (conjuncts c) else tab c |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
463 |
| c::cs => conjI OF [prove_conj tab [c], prove_conj tab cs]; |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
464 |
|
41453 | 465 |
fun conj_ac_rule eq = |
466 |
let |
|
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
467 |
val (l,r) = Thm.dest_equals eq |
46497
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
wenzelm
parents:
42793
diff
changeset
|
468 |
val ctabl = mk_conj_tab (Thm.assume (Thm.apply @{cterm Trueprop} l)) |
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
wenzelm
parents:
42793
diff
changeset
|
469 |
val ctabr = mk_conj_tab (Thm.assume (Thm.apply @{cterm Trueprop} r)) |
59582 | 470 |
fun tabl c = the (Termtab.lookup ctabl (Thm.term_of c)) |
471 |
fun tabr c = the (Termtab.lookup ctabr (Thm.term_of c)) |
|
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
472 |
val thl = prove_conj tabl (conjuncts r) |> implies_intr_hyps |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
473 |
val thr = prove_conj tabr (conjuncts l) |> implies_intr_hyps |
60801 | 474 |
val eqI = Thm.instantiate' [] [SOME l, SOME r] @{thm iffI} |
41453 | 475 |
in Thm.implies_elim (Thm.implies_elim eqI thl) thr |> mk_meta_eq end; |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
476 |
|
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
477 |
(* END FIXME.*) |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
478 |
|
41453 | 479 |
(* Conversion for the equivalence of existential statements where |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
480 |
EX quantifiers are rearranged differently *) |
61075 | 481 |
fun ext ctxt T = Thm.cterm_of ctxt (Const (@{const_name Ex}, (T --> @{typ bool}) --> @{typ bool})) |
482 |
fun mk_ex ctxt v t = Thm.apply (ext ctxt (Thm.typ_of_cterm v)) (Thm.lambda v t) |
|
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
483 |
|
59582 | 484 |
fun choose v th th' = case Thm.concl_of th of |
41453 | 485 |
@{term Trueprop} $ (Const(@{const_name Ex},_)$_) => |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
486 |
let |
59582 | 487 |
val p = (funpow 2 Thm.dest_arg o Thm.cprop_of) th |
59586 | 488 |
val T = (hd o Thm.dest_ctyp o Thm.ctyp_of_cterm) p |
41453 | 489 |
val th0 = Conv.fconv_rule (Thm.beta_conversion true) |
60801 | 490 |
(Thm.instantiate' [SOME T] [SOME p, (SOME o Thm.dest_arg o Thm.cprop_of) th'] exE) |
41453 | 491 |
val pv = (Thm.rhs_of o Thm.beta_conversion true) |
46497
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
wenzelm
parents:
42793
diff
changeset
|
492 |
(Thm.apply @{cterm Trueprop} (Thm.apply p v)) |
41453 | 493 |
val th1 = Thm.forall_intr v (Thm.implies_intr pv th') |
494 |
in Thm.implies_elim (Thm.implies_elim th0 th) th1 end |
|
495 |
| _ => error "" (* FIXME ? *) |
|
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
496 |
|
61075 | 497 |
fun simple_choose ctxt v th = |
498 |
choose v (Thm.assume ((Thm.apply @{cterm Trueprop} o mk_ex ctxt v) |
|
60949 | 499 |
(Thm.dest_arg (hd (Thm.chyps_of th))))) th |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
500 |
|
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
501 |
|
41453 | 502 |
fun mkexi v th = |
503 |
let |
|
46497
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
wenzelm
parents:
42793
diff
changeset
|
504 |
val p = Thm.lambda v (Thm.dest_arg (Thm.cprop_of th)) |
41453 | 505 |
in Thm.implies_elim |
506 |
(Conv.fconv_rule (Thm.beta_conversion true) |
|
60801 | 507 |
(Thm.instantiate' [SOME (Thm.ctyp_of_cterm v)] [SOME p, SOME v] @{thm exI})) |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
508 |
th |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
509 |
end |
61075 | 510 |
fun ex_eq_conv ctxt t = |
41453 | 511 |
let |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
512 |
val (p0,q0) = Thm.dest_binop t |
41453 | 513 |
val (vs',P) = strip_exists p0 |
514 |
val (vs,_) = strip_exists q0 |
|
46497
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
wenzelm
parents:
42793
diff
changeset
|
515 |
val th = Thm.assume (Thm.apply @{cterm Trueprop} P) |
61075 | 516 |
val th1 = implies_intr_hyps (fold (simple_choose ctxt) vs' (fold mkexi vs th)) |
517 |
val th2 = implies_intr_hyps (fold (simple_choose ctxt) vs (fold mkexi vs' th)) |
|
59582 | 518 |
val p = (Thm.dest_arg o Thm.dest_arg1 o Thm.cprop_of) th1 |
519 |
val q = (Thm.dest_arg o Thm.dest_arg o Thm.cprop_of) th1 |
|
60801 | 520 |
in Thm.implies_elim (Thm.implies_elim (Thm.instantiate' [] [SOME p, SOME q] iffI) th1) th2 |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
521 |
|> mk_meta_eq |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
522 |
end; |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
523 |
|
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
524 |
|
59582 | 525 |
fun getname v = case Thm.term_of v of |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
526 |
Free(s,_) => s |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
527 |
| Var ((s,_),_) => s |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
528 |
| _ => "x" |
46497
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
wenzelm
parents:
42793
diff
changeset
|
529 |
fun mk_eq s t = Thm.apply (Thm.apply @{cterm "op == :: bool => _"} s) t |
61075 | 530 |
fun mk_exists ctxt v th = Drule.arg_cong_rule (ext ctxt (Thm.typ_of_cterm v)) |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
531 |
(Thm.abstract_rule (getname v) v th) |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47432
diff
changeset
|
532 |
fun simp_ex_conv ctxt = |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47432
diff
changeset
|
533 |
Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps @{thms simp_thms(39)}) |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
534 |
|
60818 | 535 |
fun frees t = Drule.cterm_add_frees t []; |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
536 |
fun free_in v t = member op aconvc (frees t) v; |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
537 |
|
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
538 |
val vsubst = let |
41453 | 539 |
fun vsubst (t,v) tm = |
46497
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
wenzelm
parents:
42793
diff
changeset
|
540 |
(Thm.rhs_of o Thm.beta_conversion false) (Thm.apply (Thm.lambda v tm) t) |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
541 |
in fold vsubst end; |
23252 | 542 |
|
543 |
||
544 |
(** main **) |
|
545 |
||
546 |
fun ring_and_ideal_conv |
|
54251 | 547 |
{vars = _, semiring = (sr_ops, _), ring = (r_ops, _), |
548 |
field = (f_ops, _), idom, ideal} |
|
23252 | 549 |
dest_const mk_const ring_eq_conv ring_normalize_conv = |
550 |
let |
|
551 |
val [add_pat, mul_pat, pow_pat, zero_tm, one_tm] = sr_ops; |
|
552 |
val [ring_add_tm, ring_mul_tm, ring_pow_tm] = |
|
41453 | 553 |
map Thm.dest_fun2 [add_pat, mul_pat, pow_pat]; |
23252 | 554 |
|
555 |
val (ring_sub_tm, ring_neg_tm) = |
|
556 |
(case r_ops of |
|
41453 | 557 |
[sub_pat, neg_pat] => (Thm.dest_fun2 sub_pat, Thm.dest_fun neg_pat) |
30866 | 558 |
|_ => (@{cterm "True"}, @{cterm "True"})); |
559 |
||
560 |
val (field_div_tm, field_inv_tm) = |
|
561 |
(case f_ops of |
|
41453 | 562 |
[div_pat, inv_pat] => (Thm.dest_fun2 div_pat, Thm.dest_fun inv_pat) |
30866 | 563 |
| _ => (@{cterm "True"}, @{cterm "True"})); |
23252 | 564 |
|
565 |
val [idom_thm, neq_thm] = idom; |
|
41453 | 566 |
val [idl_sub, idl_add0] = |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
567 |
if length ideal = 2 then ideal else [eq_commute, eq_commute] |
30866 | 568 |
fun ring_dest_neg t = |
41453 | 569 |
let val (l,r) = Thm.dest_comb t |
59582 | 570 |
in if Term.could_unify(Thm.term_of l, Thm.term_of ring_neg_tm) then r |
30866 | 571 |
else raise CTERM ("ring_dest_neg", [t]) |
572 |
end |
|
23252 | 573 |
|
30866 | 574 |
fun field_dest_inv t = |
41453 | 575 |
let val (l,r) = Thm.dest_comb t in |
59582 | 576 |
if Term.could_unify (Thm.term_of l, Thm.term_of field_inv_tm) then r |
30866 | 577 |
else raise CTERM ("field_dest_inv", [t]) |
23252 | 578 |
end |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
579 |
val ring_dest_add = dest_binary ring_add_tm; |
23252 | 580 |
val ring_mk_add = mk_binop ring_add_tm; |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
581 |
val ring_dest_sub = dest_binary ring_sub_tm; |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
582 |
val ring_dest_mul = dest_binary ring_mul_tm; |
23252 | 583 |
val ring_mk_mul = mk_binop ring_mul_tm; |
30866 | 584 |
val field_dest_div = dest_binary field_div_tm; |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
585 |
val ring_dest_pow = dest_binary ring_pow_tm; |
23252 | 586 |
val ring_mk_pow = mk_binop ring_pow_tm ; |
587 |
fun grobvars tm acc = |
|
588 |
if can dest_const tm then acc |
|
41453 | 589 |
else if can ring_dest_neg tm then grobvars (Thm.dest_arg tm) acc |
590 |
else if can ring_dest_pow tm then grobvars (Thm.dest_arg1 tm) acc |
|
23252 | 591 |
else if can ring_dest_add tm orelse can ring_dest_sub tm |
592 |
orelse can ring_dest_mul tm |
|
41453 | 593 |
then grobvars (Thm.dest_arg1 tm) (grobvars (Thm.dest_arg tm) acc) |
30866 | 594 |
else if can field_dest_inv tm |
595 |
then |
|
41453 | 596 |
let val gvs = grobvars (Thm.dest_arg tm) [] |
30866 | 597 |
in if null gvs then acc else tm::acc |
598 |
end |
|
599 |
else if can field_dest_div tm then |
|
41453 | 600 |
let val lvs = grobvars (Thm.dest_arg1 tm) acc |
601 |
val gvs = grobvars (Thm.dest_arg tm) [] |
|
30866 | 602 |
in if null gvs then lvs else tm::acc |
41453 | 603 |
end |
23252 | 604 |
else tm::acc ; |
605 |
||
606 |
fun grobify_term vars tm = |
|
607 |
((if not (member (op aconvc) vars tm) then raise CTERM ("Not a variable", [tm]) else |
|
608 |
[(rat_1,map (fn i => if i aconvc tm then 1 else 0) vars)]) |
|
609 |
handle CTERM _ => |
|
610 |
((let val x = dest_const tm |
|
54251 | 611 |
in if x =/ rat_0 then [] else [(x,map (K 0) vars)] |
23252 | 612 |
end) |
613 |
handle ERROR _ => |
|
614 |
((grob_neg(grobify_term vars (ring_dest_neg tm))) |
|
615 |
handle CTERM _ => |
|
616 |
( |
|
30866 | 617 |
(grob_inv(grobify_term vars (field_dest_inv tm))) |
41453 | 618 |
handle CTERM _ => |
23252 | 619 |
((let val (l,r) = ring_dest_add tm |
620 |
in grob_add (grobify_term vars l) (grobify_term vars r) |
|
621 |
end) |
|
622 |
handle CTERM _ => |
|
623 |
((let val (l,r) = ring_dest_sub tm |
|
624 |
in grob_sub (grobify_term vars l) (grobify_term vars r) |
|
625 |
end) |
|
626 |
handle CTERM _ => |
|
627 |
((let val (l,r) = ring_dest_mul tm |
|
628 |
in grob_mul (grobify_term vars l) (grobify_term vars r) |
|
629 |
end) |
|
630 |
handle CTERM _ => |
|
30866 | 631 |
( (let val (l,r) = field_dest_div tm |
23252 | 632 |
in grob_div (grobify_term vars l) (grobify_term vars r) |
633 |
end) |
|
30866 | 634 |
handle CTERM _ => |
23252 | 635 |
((let val (l,r) = ring_dest_pow tm |
59582 | 636 |
in grob_pow vars (grobify_term vars l) ((Thm.term_of #> HOLogic.dest_number #> snd) r) |
23252 | 637 |
end) |
638 |
handle CTERM _ => error "grobify_term: unknown or invalid term"))))))))); |
|
41453 | 639 |
val eq_tm = idom_thm |> concl |> Thm.dest_arg |> Thm.dest_arg |> Thm.dest_fun2; |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
640 |
val dest_eq = dest_binary eq_tm; |
23252 | 641 |
|
642 |
fun grobify_equation vars tm = |
|
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
643 |
let val (l,r) = dest_binary eq_tm tm |
23252 | 644 |
in grob_sub (grobify_term vars l) (grobify_term vars r) |
645 |
end; |
|
646 |
||
647 |
fun grobify_equations tm = |
|
648 |
let |
|
649 |
val cjs = conjs tm |
|
41453 | 650 |
val rawvars = |
651 |
fold_rev (fn eq => fn a => grobvars (Thm.dest_arg1 eq) (grobvars (Thm.dest_arg eq) a)) cjs [] |
|
59582 | 652 |
val vars = sort (fn (x, y) => Term_Ord.term_ord (Thm.term_of x, Thm.term_of y)) |
23252 | 653 |
(distinct (op aconvc) rawvars) |
654 |
in (vars,map (grobify_equation vars) cjs) |
|
655 |
end; |
|
656 |
||
657 |
val holify_polynomial = |
|
658 |
let fun holify_varpow (v,n) = |
|
37388 | 659 |
if n = 1 then v else ring_mk_pow v (Numeral.mk_cnumber @{ctyp nat} n) (* FIXME *) |
23252 | 660 |
fun holify_monomial vars (c,m) = |
24630
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
wenzelm
parents:
23579
diff
changeset
|
661 |
let val xps = map holify_varpow (filter (fn (_,n) => n <> 0) (vars ~~ m)) |
23252 | 662 |
in end_itlist ring_mk_mul (mk_const c :: xps) |
663 |
end |
|
664 |
fun holify_polynomial vars p = |
|
23579 | 665 |
if null p then mk_const (rat_0) |
23252 | 666 |
else end_itlist ring_mk_add (map (holify_monomial vars) p) |
667 |
in holify_polynomial |
|
668 |
end ; |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47432
diff
changeset
|
669 |
|
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47432
diff
changeset
|
670 |
fun idom_rule ctxt = simplify (put_simpset HOL_basic_ss ctxt addsimps [idom_thm]); |
23252 | 671 |
fun prove_nz n = eqF_elim |
672 |
(ring_eq_conv(mk_binop eq_tm (mk_const n) (mk_const(rat_0)))); |
|
673 |
val neq_01 = prove_nz (rat_1); |
|
674 |
fun neq_rule n th = [prove_nz n, th] MRS neq_thm; |
|
41453 | 675 |
fun mk_add th1 = Thm.combination (Drule.arg_cong_rule ring_add_tm th1); |
23252 | 676 |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47432
diff
changeset
|
677 |
fun refute ctxt tm = |
23557 | 678 |
if tm aconvc false_tm then assume_Trueprop tm else |
29800
f73a68c9d810
Now catch ERROR exception thrown by find and friends
chaieb
parents:
29269
diff
changeset
|
679 |
((let |
59970 | 680 |
val (nths0,eths0) = List.partition (is_neg o concl) (HOLogic.conj_elims ctxt (assume_Trueprop tm)) |
41453 | 681 |
val nths = filter (is_eq o Thm.dest_arg o concl) nths0 |
23252 | 682 |
val eths = filter (is_eq o concl) eths0 |
683 |
in |
|
684 |
if null eths then |
|
685 |
let |
|
59970 | 686 |
val th1 = end_itlist (fn th1 => fn th2 => idom_rule ctxt (HOLogic.conj_intr ctxt th1 th2)) nths |
41453 | 687 |
val th2 = |
688 |
Conv.fconv_rule |
|
689 |
((Conv.arg_conv #> Conv.arg_conv) (Conv.binop_conv ring_normalize_conv)) th1 |
|
690 |
val conc = th2 |> concl |> Thm.dest_arg |
|
54251 | 691 |
val (l,_) = conc |> dest_eq |
46497
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
wenzelm
parents:
42793
diff
changeset
|
692 |
in Thm.implies_intr (Thm.apply cTrp tm) |
41453 | 693 |
(Thm.equal_elim (Drule.arg_cong_rule cTrp (eqF_intr th2)) |
694 |
(Thm.reflexive l |> mk_object_eq)) |
|
23252 | 695 |
end |
696 |
else |
|
697 |
let |
|
698 |
val (vars,l,cert,noteqth) =( |
|
699 |
if null nths then |
|
700 |
let val (vars,pols) = grobify_equations(list_mk_conj(map concl eths)) |
|
701 |
val (l,cert) = grobner_weak vars pols |
|
702 |
in (vars,l,cert,neq_01) |
|
703 |
end |
|
704 |
else |
|
705 |
let |
|
59970 | 706 |
val nth = end_itlist (fn th1 => fn th2 => idom_rule ctxt (HOLogic.conj_intr ctxt th1 th2)) nths |
23252 | 707 |
val (vars,pol::pols) = |
41453 | 708 |
grobify_equations(list_mk_conj(Thm.dest_arg(concl nth)::map concl eths)) |
23252 | 709 |
val (deg,l,cert) = grobner_strong vars pols pol |
41453 | 710 |
val th1 = |
711 |
Conv.fconv_rule ((Conv.arg_conv o Conv.arg_conv) (Conv.binop_conv ring_normalize_conv)) nth |
|
59970 | 712 |
val th2 = funpow deg (idom_rule ctxt o HOLogic.conj_intr ctxt th1) neq_01 |
23252 | 713 |
in (vars,l,cert,th2) |
714 |
end) |
|
54251 | 715 |
val cert_pos = map (fn (i,p) => (i,filter (fn (c,_) => c >/ rat_0) p)) cert |
23252 | 716 |
val cert_neg = map (fn (i,p) => (i,map (fn (c,m) => (minus_rat c,m)) |
54251 | 717 |
(filter (fn (c,_) => c </ rat_0) p))) cert |
23252 | 718 |
val herts_pos = map (fn (i,p) => (i,holify_polynomial vars p)) cert_pos |
719 |
val herts_neg = map (fn (i,p) => (i,holify_polynomial vars p)) cert_neg |
|
720 |
fun thm_fn pols = |
|
41453 | 721 |
if null pols then Thm.reflexive(mk_const rat_0) else |
23252 | 722 |
end_itlist mk_add |
46497
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
wenzelm
parents:
42793
diff
changeset
|
723 |
(map (fn (i,p) => Drule.arg_cong_rule (Thm.apply ring_mul_tm p) |
24630
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
wenzelm
parents:
23579
diff
changeset
|
724 |
(nth eths i |> mk_meta_eq)) pols) |
23252 | 725 |
val th1 = thm_fn herts_pos |
726 |
val th2 = thm_fn herts_neg |
|
59970 | 727 |
val th3 = HOLogic.conj_intr ctxt (mk_add (Thm.symmetric th1) th2 |> mk_object_eq) noteqth |
41453 | 728 |
val th4 = |
729 |
Conv.fconv_rule ((Conv.arg_conv o Conv.arg_conv o Conv.binop_conv) ring_normalize_conv) |
|
730 |
(neq_rule l th3) |
|
54251 | 731 |
val (l, _) = dest_eq(Thm.dest_arg(concl th4)) |
46497
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
wenzelm
parents:
42793
diff
changeset
|
732 |
in Thm.implies_intr (Thm.apply cTrp tm) |
41453 | 733 |
(Thm.equal_elim (Drule.arg_cong_rule cTrp (eqF_intr th4)) |
734 |
(Thm.reflexive l |> mk_object_eq)) |
|
23252 | 735 |
end |
41453 | 736 |
end) handle ERROR _ => raise CTERM ("Groebner-refute: unable to refute",[tm])) |
23252 | 737 |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47432
diff
changeset
|
738 |
fun ring ctxt tm = |
23252 | 739 |
let |
740 |
fun mk_forall x p = |
|
61075 | 741 |
let |
742 |
val T = Thm.typ_of_cterm x; |
|
743 |
val all = Thm.cterm_of ctxt (Const (@{const_name All}, (T --> @{typ bool}) --> @{typ bool})) |
|
744 |
in Thm.apply all (Thm.lambda x p) end |
|
60818 | 745 |
val avs = Drule.cterm_add_frees tm [] |
23252 | 746 |
val P' = fold mk_forall avs tm |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47432
diff
changeset
|
747 |
val th1 = initial_conv ctxt (mk_neg P') |
23252 | 748 |
val (evs,bod) = strip_exists(concl th1) in |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
749 |
if is_forall bod then raise CTERM("ring: non-universal formula",[tm]) |
23252 | 750 |
else |
751 |
let |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47432
diff
changeset
|
752 |
val th1a = weak_dnf_conv ctxt bod |
23252 | 753 |
val boda = concl th1a |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47432
diff
changeset
|
754 |
val th2a = refute_disj (refute ctxt) boda |
23252 | 755 |
val th2b = [mk_object_eq th1a, (th2a COMP notI) COMP PFalse'] MRS trans |
41453 | 756 |
val th2 = fold (fn v => fn th => (Thm.forall_intr v th) COMP allI) evs (th2b RS PFalse) |
757 |
val th3 = |
|
758 |
Thm.equal_elim |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47432
diff
changeset
|
759 |
(Simplifier.rewrite (put_simpset HOL_basic_ss ctxt addsimps [not_ex RS sym]) |
59582 | 760 |
(th2 |> Thm.cprop_of)) th2 |
23252 | 761 |
in specl avs |
762 |
([[[mk_object_eq th1, th3 RS PFalse'] MRS trans] MRS PFalse] MRS notnotD) |
|
763 |
end |
|
764 |
end |
|
765 |
fun ideal tms tm ord = |
|
766 |
let |
|
767 |
val rawvars = fold_rev grobvars (tm::tms) [] |
|
59582 | 768 |
val vars = sort ord (distinct (fn (x,y) => (Thm.term_of x) aconv (Thm.term_of y)) rawvars) |
23252 | 769 |
val pols = map (grobify_term vars) tms |
770 |
val pol = grobify_term vars tm |
|
771 |
val cert = grobner_ideal vars pols pol |
|
33063 | 772 |
in map_range (fn n => these (AList.lookup (op =) cert n) |> holify_polynomial vars) |
773 |
(length pols) |
|
23252 | 774 |
end |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
775 |
|
41453 | 776 |
fun poly_eq_conv t = |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
777 |
let val (a,b) = Thm.dest_binop t |
41453 | 778 |
in Conv.fconv_rule (Conv.arg_conv (Conv.arg1_conv ring_normalize_conv)) |
60801 | 779 |
(Thm.instantiate' [] [SOME a, SOME b] idl_sub) |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
780 |
end |
61144 | 781 |
|
782 |
val poly_eq_simproc = |
|
41453 | 783 |
let |
61144 | 784 |
fun proc ct = |
785 |
let val th = poly_eq_conv ct |
|
786 |
in if Thm.is_reflexive th then NONE else SOME th end |
|
787 |
in |
|
788 |
Simplifier.cert_simproc (Thm.theory_of_thm idl_sub) "poly_eq_simproc" |
|
789 |
{lhss = [Thm.term_of (Thm.lhs_of idl_sub)], |
|
790 |
proc = fn _ => fn _ => proc, identifier = []} |
|
791 |
end; |
|
792 |
||
793 |
val poly_eq_ss = |
|
794 |
simpset_of (put_simpset HOL_basic_ss @{context} |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47432
diff
changeset
|
795 |
addsimps @{thms simp_thms} |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47432
diff
changeset
|
796 |
addsimprocs [poly_eq_simproc]) |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
797 |
|
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
798 |
local |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
799 |
fun is_defined v t = |
41453 | 800 |
let |
801 |
val mons = striplist(dest_binary ring_add_tm) t |
|
802 |
in member (op aconvc) mons v andalso |
|
803 |
forall (fn m => v aconvc m |
|
60818 | 804 |
orelse not(member (op aconvc) (Drule.cterm_add_frees m []) v)) mons |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
805 |
end |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
806 |
|
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
807 |
fun isolate_variable vars tm = |
41453 | 808 |
let |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
809 |
val th = poly_eq_conv tm |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
810 |
val th' = (sym_conv then_conv poly_eq_conv) tm |
41453 | 811 |
val (v,th1) = |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
812 |
case find_first(fn v=> is_defined v (Thm.dest_arg1 (Thm.rhs_of th))) vars of |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
813 |
SOME v => (v,th') |
41453 | 814 |
| NONE => (the (find_first |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
815 |
(fn v => is_defined v (Thm.dest_arg1 (Thm.rhs_of th'))) vars) ,th) |
41453 | 816 |
val th2 = Thm.transitive th1 |
60801 | 817 |
(Thm.instantiate' [] [(SOME o Thm.dest_arg1 o Thm.rhs_of) th1, SOME v] |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
818 |
idl_add0) |
41453 | 819 |
in Conv.fconv_rule(funpow 2 Conv.arg_conv ring_normalize_conv) th2 |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
820 |
end |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
821 |
in |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47432
diff
changeset
|
822 |
fun unwind_polys_conv ctxt tm = |
41453 | 823 |
let |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
824 |
val (vars,bod) = strip_exists tm |
38795
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents:
38786
diff
changeset
|
825 |
val cjs = striplist (dest_binary @{cterm HOL.conj}) bod |
41453 | 826 |
val th1 = (the (get_first (try (isolate_variable vars)) cjs) |
51930
52fd62618631
prefer explicitly qualified exceptions, which is particular important for robust handlers;
wenzelm
parents:
51717
diff
changeset
|
827 |
handle Option.Option => raise CTERM ("unwind_polys_conv",[tm])) |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
828 |
val eq = Thm.lhs_of th1 |
38795
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents:
38786
diff
changeset
|
829 |
val bod' = list_mk_binop @{cterm HOL.conj} (eq::(remove op aconvc eq cjs)) |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
830 |
val th2 = conj_ac_rule (mk_eq bod bod') |
41453 | 831 |
val th3 = |
832 |
Thm.transitive th2 |
|
833 |
(Drule.binop_cong_rule @{cterm HOL.conj} th1 |
|
834 |
(Thm.reflexive (Thm.dest_arg (Thm.rhs_of th2)))) |
|
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
835 |
val v = Thm.dest_arg1(Thm.dest_arg1(Thm.rhs_of th3)) |
61075 | 836 |
val th4 = Conv.fconv_rule (Conv.arg_conv (simp_ex_conv ctxt)) (mk_exists ctxt v th3) |
837 |
val th5 = ex_eq_conv ctxt (mk_eq tm (fold (mk_ex ctxt) (remove op aconvc v vars) (Thm.lhs_of th4))) |
|
838 |
in Thm.transitive th5 (fold (mk_exists ctxt) (remove op aconvc v vars) th4) |
|
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
839 |
end; |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
840 |
end |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
841 |
|
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
842 |
local |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
843 |
fun scrub_var v m = |
41453 | 844 |
let |
845 |
val ps = striplist ring_dest_mul m |
|
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
846 |
val ps' = remove op aconvc v ps |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
847 |
in if null ps' then one_tm else fold1 ring_mk_mul ps' |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
848 |
end |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
849 |
fun find_multipliers v mons = |
41453 | 850 |
let |
851 |
val mons1 = filter (fn m => free_in v m) mons |
|
852 |
val mons2 = map (scrub_var v) mons1 |
|
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
853 |
in if null mons2 then zero_tm else fold1 ring_mk_add mons2 |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
854 |
end |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
855 |
|
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
856 |
fun isolate_monomials vars tm = |
41453 | 857 |
let |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
858 |
val (cmons,vmons) = |
33049
c38f02fdf35d
curried inter as canonical list operation (beware of argument order)
haftmann
parents:
33039
diff
changeset
|
859 |
List.partition (fn m => null (inter (op aconvc) vars (frees m))) |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
860 |
(striplist ring_dest_add tm) |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
861 |
val cofactors = map (fn v => find_multipliers v vmons) vars |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
862 |
val cnc = if null cmons then zero_tm |
46497
89ccf66aa73d
renamed Thm.capply to Thm.apply, and Thm.cabs to Thm.lambda in conformance with similar operations in structure Term and Logic;
wenzelm
parents:
42793
diff
changeset
|
863 |
else Thm.apply ring_neg_tm |
41453 | 864 |
(list_mk_binop ring_add_tm cmons) |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
865 |
in (cofactors,cnc) |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
866 |
end; |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
867 |
|
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
868 |
fun isolate_variables evs ps eq = |
41453 | 869 |
let |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
870 |
val vars = filter (fn v => free_in v eq) evs |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
871 |
val (qs,p) = isolate_monomials vars eq |
41453 | 872 |
val rs = ideal (qs @ ps) p |
59582 | 873 |
(fn (s,t) => Term_Ord.term_ord (Thm.term_of s, Thm.term_of t)) |
33957 | 874 |
in (eq, take (length qs) rs ~~ vars) |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
875 |
end; |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
876 |
fun subst_in_poly i p = Thm.rhs_of (ring_normalize_conv (vsubst i p)); |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
877 |
in |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
878 |
fun solve_idealism evs ps eqs = |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
879 |
if null evs then [] else |
41453 | 880 |
let |
33035 | 881 |
val (eq,cfs) = get_first (try (isolate_variables evs ps)) eqs |> the |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
882 |
val evs' = subtract op aconvc evs (map snd cfs) |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
883 |
val eqs' = map (subst_in_poly cfs) (remove op aconvc eq eqs) |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
884 |
in cfs @ solve_idealism evs' ps eqs' |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
885 |
end; |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
886 |
end; |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
887 |
|
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
888 |
|
41453 | 889 |
in {ring_conv = ring, simple_ideal = ideal, multi_ideal = solve_idealism, |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
890 |
poly_eq_ss = poly_eq_ss, unwind_conv = unwind_polys_conv} |
23252 | 891 |
end; |
892 |
||
893 |
||
894 |
fun find_term bounds tm = |
|
59582 | 895 |
(case Thm.term_of tm of |
38864
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
haftmann
parents:
38795
diff
changeset
|
896 |
Const (@{const_name HOL.eq}, T) $ _ $ _ => |
23252 | 897 |
if domain_type T = HOLogic.boolT then find_args bounds tm |
41453 | 898 |
else Thm.dest_arg tm |
899 |
| Const (@{const_name Not}, _) $ _ => find_term bounds (Thm.dest_arg tm) |
|
900 |
| Const (@{const_name All}, _) $ _ => find_body bounds (Thm.dest_arg tm) |
|
901 |
| Const (@{const_name Ex}, _) $ _ => find_body bounds (Thm.dest_arg tm) |
|
38795
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents:
38786
diff
changeset
|
902 |
| Const (@{const_name HOL.conj}, _) $ _ $ _ => find_args bounds tm |
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents:
38786
diff
changeset
|
903 |
| Const (@{const_name HOL.disj}, _) $ _ $ _ => find_args bounds tm |
38786
e46e7a9cb622
formerly unnamed infix impliciation now named HOL.implies
haftmann
parents:
38558
diff
changeset
|
904 |
| Const (@{const_name HOL.implies}, _) $ _ $ _ => find_args bounds tm |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
905 |
| @{term "op ==>"} $_$_ => find_args bounds tm |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
906 |
| Const("op ==",_)$_$_ => find_args bounds tm |
41453 | 907 |
| @{term Trueprop}$_ => find_term bounds (Thm.dest_arg tm) |
23252 | 908 |
| _ => raise TERM ("find_term", [])) |
909 |
and find_args bounds tm = |
|
910 |
let val (t, u) = Thm.dest_binop tm |
|
911 |
in (find_term bounds t handle TERM _ => find_term bounds u) end |
|
912 |
and find_body bounds b = |
|
41453 | 913 |
let val (_, b') = Thm.dest_abs (SOME (Name.bound bounds)) b |
23252 | 914 |
in find_term (bounds + 1) b' end; |
915 |
||
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
916 |
|
41453 | 917 |
fun get_ring_ideal_convs ctxt form = |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
918 |
case try (find_term 0) form of |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
919 |
NONE => NONE |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
920 |
| SOME tm => |
36752
cf558aeb35b0
delete Groebner_Basis directory -- only one file left
haftmann
parents:
36723
diff
changeset
|
921 |
(case Semiring_Normalizer.match ctxt tm of |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
922 |
NONE => NONE |
54251 | 923 |
| SOME (res as (theory, {is_const = _, dest_const, |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
924 |
mk_const, conv = ring_eq_conv})) => |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
925 |
SOME (ring_and_ideal_conv theory |
59586 | 926 |
dest_const (mk_const (Thm.ctyp_of_cterm tm)) (ring_eq_conv ctxt) |
36752
cf558aeb35b0
delete Groebner_Basis directory -- only one file left
haftmann
parents:
36723
diff
changeset
|
927 |
(Semiring_Normalizer.semiring_normalize_wrapper ctxt res))) |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
928 |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47432
diff
changeset
|
929 |
fun presimplify ctxt add_thms del_thms = |
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47432
diff
changeset
|
930 |
asm_full_simp_tac (put_simpset HOL_basic_ss ctxt |
57951 | 931 |
addsimps (Named_Theorems.get ctxt @{named_theorems algebra}) |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47432
diff
changeset
|
932 |
delsimps del_thms addsimps add_thms); |
36702
b455ebd63799
moved presimplification rules for algebraic methods into named thms functor
haftmann
parents:
36700
diff
changeset
|
933 |
|
23579 | 934 |
fun ring_tac add_ths del_ths ctxt = |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
54251
diff
changeset
|
935 |
Object_Logic.full_atomize_tac ctxt |
36702
b455ebd63799
moved presimplification rules for algebraic methods into named thms functor
haftmann
parents:
36700
diff
changeset
|
936 |
THEN' presimplify ctxt add_ths del_ths |
23579 | 937 |
THEN' CSUBGOAL (fn (p, i) => |
60752 | 938 |
resolve_tac ctxt [let val form = Object_Logic.dest_judgment ctxt p |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
939 |
in case get_ring_ideal_convs ctxt form of |
41453 | 940 |
NONE => Thm.reflexive form |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47432
diff
changeset
|
941 |
| SOME thy => #ring_conv thy ctxt form |
60752 | 942 |
end] i |
23579 | 943 |
handle TERM _ => no_tac |
944 |
| CTERM _ => no_tac |
|
945 |
| THM _ => no_tac); |
|
23334 | 946 |
|
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
947 |
local |
59582 | 948 |
fun lhs t = case Thm.term_of t of |
38864
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
haftmann
parents:
38795
diff
changeset
|
949 |
Const(@{const_name HOL.eq},_)$_$_ => Thm.dest_arg1 t |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
950 |
| _=> raise CTERM ("ideal_tac - lhs",[t]) |
60752 | 951 |
fun exitac _ NONE = no_tac |
952 |
| exitac ctxt (SOME y) = |
|
60801 | 953 |
resolve_tac ctxt [Thm.instantiate' [SOME (Thm.ctyp_of_cterm y)] [NONE,SOME y] exI] 1 |
52086 | 954 |
|
955 |
val claset = claset_of @{context} |
|
41453 | 956 |
in |
957 |
fun ideal_tac add_ths del_ths ctxt = |
|
36702
b455ebd63799
moved presimplification rules for algebraic methods into named thms functor
haftmann
parents:
36700
diff
changeset
|
958 |
presimplify ctxt add_ths del_ths |
27671 | 959 |
THEN' |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
960 |
CSUBGOAL (fn (p, i) => |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
961 |
case get_ring_ideal_convs ctxt p of |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
962 |
NONE => no_tac |
41453 | 963 |
| SOME thy => |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
964 |
let |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
965 |
fun poly_exists_tac {asms = asms, concl = concl, prems = prems, |
54251 | 966 |
params = _, context = ctxt, schematics = _} = |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
967 |
let |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
968 |
val (evs,bod) = strip_exists (Thm.dest_arg concl) |
41453 | 969 |
val ps = map_filter (try (lhs o Thm.dest_arg)) asms |
970 |
val cfs = (map swap o #multi_ideal thy evs ps) |
|
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
971 |
(map Thm.dest_arg1 (conjuncts bod)) |
60752 | 972 |
val ws = map (exitac ctxt o AList.lookup op aconvc cfs) evs |
41453 | 973 |
in EVERY (rev ws) THEN Method.insert_tac prems 1 |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
974 |
THEN ring_tac add_ths del_ths ctxt 1 |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
975 |
end |
41453 | 976 |
in |
52086 | 977 |
clarify_tac (put_claset claset ctxt) i |
54742
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
wenzelm
parents:
54251
diff
changeset
|
978 |
THEN Object_Logic.full_atomize_tac ctxt i |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47432
diff
changeset
|
979 |
THEN asm_full_simp_tac (put_simpset (#poly_eq_ss thy) ctxt) i |
52086 | 980 |
THEN clarify_tac (put_claset claset ctxt) i |
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
47432
diff
changeset
|
981 |
THEN (REPEAT (CONVERSION (#unwind_conv thy ctxt) i)) |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
982 |
THEN SUBPROOF poly_exists_tac ctxt i |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
983 |
end |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
984 |
handle TERM _ => no_tac |
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
985 |
| CTERM _ => no_tac |
41453 | 986 |
| THM _ => no_tac); |
23252 | 987 |
end; |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
988 |
|
41453 | 989 |
fun algebra_tac add_ths del_ths ctxt i = |
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
990 |
ring_tac add_ths del_ths ctxt i ORELSE ideal_tac add_ths del_ths ctxt i |
41453 | 991 |
|
25251
759bffe1d416
(1) signatures updated according to normalizer_data.ML (added field ideal in entry);
chaieb
parents:
24913
diff
changeset
|
992 |
end; |