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