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