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