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