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