src/HOL/Tools/Groebner_Basis/normalizer_data.ML
author bulwahn
Thu, 12 Nov 2009 20:38:57 +0100
changeset 33649 854173fcd21c
parent 33519 e31a85f92ce9
child 35624 c4e29a0bb8c1
permissions -rw-r--r--
added a tabled implementation of the reflexive transitive closure
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
23252
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
     1
(*  Title:      HOL/Tools/Groebner_Basis/normalizer_data.ML
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
     3
    Author:     Amine Chaieb, TU Muenchen
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
     4
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
     5
Ring normalization data.
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
     6
*)
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
     7
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
     8
signature NORMALIZER_DATA =
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
     9
sig
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    10
  type entry
23581
wenzelm
parents: 23485
diff changeset
    11
  val get: Proof.context -> simpset * (thm * entry) list
23252
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    12
  val match: Proof.context -> cterm -> entry option
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    13
  val del: attribute
30866
dd5117e2d73e now deals with devision in fields
chaieb
parents: 30722
diff changeset
    14
  val add: {semiring: cterm list * thm list, ring: cterm list * thm list, field: cterm list * thm list, idom: thm list, ideal: thm list}
23252
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    15
    -> attribute
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    16
  val funs: thm -> {is_const: morphism -> cterm -> bool,
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    17
    dest_const: morphism -> cterm -> Rat.rat,
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    18
    mk_const: morphism -> ctyp -> Rat.rat -> cterm,
24020
ed4d7abffee7 tuned signature;
wenzelm
parents: 23581
diff changeset
    19
    conv: morphism -> Proof.context -> cterm -> thm} -> declaration
23252
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    20
  val setup: theory -> theory
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    21
end;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    22
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    23
structure NormalizerData: NORMALIZER_DATA =
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    24
struct
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    25
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    26
(* data *)
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    27
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    28
type entry =
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    29
 {vars: cterm list,
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    30
  semiring: cterm list * thm list,
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    31
  ring: cterm list * thm list,
30866
dd5117e2d73e now deals with devision in fields
chaieb
parents: 30722
diff changeset
    32
  field: cterm list * thm list,
25254
0216ca99a599 Added field ideal into entry - uses by algebra method to prove the ideal membership problem
chaieb
parents: 24020
diff changeset
    33
  idom: thm list,
0216ca99a599 Added field ideal into entry - uses by algebra method to prove the ideal membership problem
chaieb
parents: 24020
diff changeset
    34
  ideal: thm list} *
23252
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    35
 {is_const: cterm -> bool,
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    36
  dest_const: cterm -> Rat.rat,
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    37
  mk_const: ctyp -> Rat.rat -> cterm,
23330
01c09922ce59 Conversion for computation on constants now depends on the context
chaieb
parents: 23260
diff changeset
    38
  conv: Proof.context -> cterm -> thm};
23252
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    39
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    40
val eq_key = Thm.eq_thm;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    41
fun eq_data arg = eq_fst eq_key arg;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    42
33519
e31a85f92ce9 adapted Generic_Data, Proof_Data;
wenzelm
parents: 30866
diff changeset
    43
structure Data = Generic_Data
23252
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    44
(
23581
wenzelm
parents: 23485
diff changeset
    45
  type T = simpset * (thm * entry) list;
23335
42b827dfa86b Changed normalizer Data to contain a simpset besides the list of instances; Before the method starts, it simplifies with this simpset; Added attribute 'algebra' to add and delete theorems from the simpset above;
chaieb
parents: 23330
diff changeset
    46
  val empty = (HOL_basic_ss, []);
23252
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    47
  val extend = I;
33519
e31a85f92ce9 adapted Generic_Data, Proof_Data;
wenzelm
parents: 30866
diff changeset
    48
  fun merge ((ss, e), (ss', e')) : T =
23581
wenzelm
parents: 23485
diff changeset
    49
    (merge_ss (ss, ss'), AList.merge eq_key (K true) (e, e'));
23252
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    50
);
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    51
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    52
val get = Data.get o Context.Proof;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    53
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    54
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    55
(* match data *)
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    56
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    57
fun match ctxt tm =
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    58
  let
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    59
    fun match_inst
25254
0216ca99a599 Added field ideal into entry - uses by algebra method to prove the ideal membership problem
chaieb
parents: 24020
diff changeset
    60
        ({vars, semiring = (sr_ops, sr_rules), 
30866
dd5117e2d73e now deals with devision in fields
chaieb
parents: 30722
diff changeset
    61
          ring = (r_ops, r_rules), field = (f_ops, f_rules), idom, ideal},
23252
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    62
         fns as {is_const, dest_const, mk_const, conv}) pat =
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    63
       let
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    64
        fun h instT =
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    65
          let
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    66
            val substT = Thm.instantiate (instT, []);
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    67
            val substT_cterm = Drule.cterm_rule substT;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    68
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    69
            val vars' = map substT_cterm vars;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    70
            val semiring' = (map substT_cterm sr_ops, map substT sr_rules);
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    71
            val ring' = (map substT_cterm r_ops, map substT r_rules);
30866
dd5117e2d73e now deals with devision in fields
chaieb
parents: 30722
diff changeset
    72
            val field' = (map substT_cterm f_ops, map substT f_rules);
23252
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    73
            val idom' = map substT idom;
25254
0216ca99a599 Added field ideal into entry - uses by algebra method to prove the ideal membership problem
chaieb
parents: 24020
diff changeset
    74
            val ideal' = map substT ideal;
23252
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    75
25254
0216ca99a599 Added field ideal into entry - uses by algebra method to prove the ideal membership problem
chaieb
parents: 24020
diff changeset
    76
            val result = ({vars = vars', semiring = semiring', 
30866
dd5117e2d73e now deals with devision in fields
chaieb
parents: 30722
diff changeset
    77
                           ring = ring', field = field', idom = idom', ideal = ideal'}, fns);
23252
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    78
          in SOME result end
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    79
      in (case try Thm.match (pat, tm) of
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    80
           NONE => NONE
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    81
         | SOME (instT, _) => h instT)
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    82
      end;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    83
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    84
    fun match_struct (_,
30866
dd5117e2d73e now deals with devision in fields
chaieb
parents: 30722
diff changeset
    85
        entry as ({semiring = (sr_ops, _), ring = (r_ops, _), field = (f_ops, _), ...}, _): entry) =
dd5117e2d73e now deals with devision in fields
chaieb
parents: 30722
diff changeset
    86
      get_first (match_inst entry) (sr_ops @ r_ops @ f_ops);
23335
42b827dfa86b Changed normalizer Data to contain a simpset besides the list of instances; Before the method starts, it simplifies with this simpset; Added attribute 'algebra' to add and delete theorems from the simpset above;
chaieb
parents: 23330
diff changeset
    87
  in get_first match_struct (snd (get ctxt)) end;
23252
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    88
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    89
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    90
(* logical content *)
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    91
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    92
val semiringN = "semiring";
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    93
val ringN = "ring";
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    94
val idomN = "idom";
25254
0216ca99a599 Added field ideal into entry - uses by algebra method to prove the ideal membership problem
chaieb
parents: 24020
diff changeset
    95
val idealN = "ideal";
30866
dd5117e2d73e now deals with devision in fields
chaieb
parents: 30722
diff changeset
    96
val fieldN = "field";
23252
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    97
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    98
fun undefined _ = raise Match;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    99
23335
42b827dfa86b Changed normalizer Data to contain a simpset besides the list of instances; Before the method starts, it simplifies with this simpset; Added attribute 'algebra' to add and delete theorems from the simpset above;
chaieb
parents: 23330
diff changeset
   100
fun del_data key = apsnd (remove eq_data (key, []));
23252
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   101
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   102
val del = Thm.declaration_attribute (Data.map o del_data);
23335
42b827dfa86b Changed normalizer Data to contain a simpset besides the list of instances; Before the method starts, it simplifies with this simpset; Added attribute 'algebra' to add and delete theorems from the simpset above;
chaieb
parents: 23330
diff changeset
   103
val add_ss = Thm.declaration_attribute 
42b827dfa86b Changed normalizer Data to contain a simpset besides the list of instances; Before the method starts, it simplifies with this simpset; Added attribute 'algebra' to add and delete theorems from the simpset above;
chaieb
parents: 23330
diff changeset
   104
   (fn th => Data.map (fn (ss,data) => (ss addsimps [th], data)));
42b827dfa86b Changed normalizer Data to contain a simpset besides the list of instances; Before the method starts, it simplifies with this simpset; Added attribute 'algebra' to add and delete theorems from the simpset above;
chaieb
parents: 23330
diff changeset
   105
42b827dfa86b Changed normalizer Data to contain a simpset besides the list of instances; Before the method starts, it simplifies with this simpset; Added attribute 'algebra' to add and delete theorems from the simpset above;
chaieb
parents: 23330
diff changeset
   106
val del_ss = Thm.declaration_attribute 
42b827dfa86b Changed normalizer Data to contain a simpset besides the list of instances; Before the method starts, it simplifies with this simpset; Added attribute 'algebra' to add and delete theorems from the simpset above;
chaieb
parents: 23330
diff changeset
   107
   (fn th => Data.map (fn (ss,data) => (ss delsimps [th], data)));
23252
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   108
30866
dd5117e2d73e now deals with devision in fields
chaieb
parents: 30722
diff changeset
   109
fun add {semiring = (sr_ops, sr_rules), ring = (r_ops, r_rules), 
dd5117e2d73e now deals with devision in fields
chaieb
parents: 30722
diff changeset
   110
         field = (f_ops, f_rules), idom, ideal} =
23252
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   111
  Thm.declaration_attribute (fn key => fn context => context |> Data.map
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   112
    let
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   113
      val ctxt = Context.proof_of context;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   114
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   115
      fun check kind name xs n =
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   116
        null xs orelse length xs = n orelse
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   117
        error ("Expected " ^ string_of_int n ^ " " ^ kind ^ " for " ^ name);
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   118
      val check_ops = check "operations";
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   119
      val check_rules = check "rules";
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   120
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   121
      val _ =
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   122
        check_ops semiringN sr_ops 5 andalso
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   123
        check_rules semiringN sr_rules 37 andalso
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   124
        check_ops ringN r_ops 2 andalso
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   125
        check_rules ringN r_rules 2 andalso
30866
dd5117e2d73e now deals with devision in fields
chaieb
parents: 30722
diff changeset
   126
        check_ops fieldN f_ops 2 andalso
dd5117e2d73e now deals with devision in fields
chaieb
parents: 30722
diff changeset
   127
        check_rules fieldN f_rules 2 andalso
23252
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   128
        check_rules idomN idom 2;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   129
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   130
      val mk_meta = LocalDefs.meta_rewrite_rule ctxt;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   131
      val sr_rules' = map mk_meta sr_rules;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   132
      val r_rules' = map mk_meta r_rules;
30866
dd5117e2d73e now deals with devision in fields
chaieb
parents: 30722
diff changeset
   133
      val f_rules' = map mk_meta f_rules;
23252
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   134
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   135
      fun rule i = nth sr_rules' (i - 1);
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   136
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   137
      val (cx, cy) = Thm.dest_binop (hd sr_ops);
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   138
      val cz = rule 34 |> Thm.rhs_of |> Thm.dest_arg |> Thm.dest_arg;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   139
      val cn = rule 36 |> Thm.rhs_of |> Thm.dest_arg |> Thm.dest_arg;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   140
      val ((clx, crx), (cly, cry)) =
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   141
        rule 13 |> Thm.rhs_of |> Thm.dest_binop |> pairself Thm.dest_binop;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   142
      val ((ca, cb), (cc, cd)) =
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   143
        rule 20 |> Thm.lhs_of |> Thm.dest_binop |> pairself Thm.dest_binop;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   144
      val cm = rule 1 |> Thm.rhs_of |> Thm.dest_arg;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   145
      val (cp, cq) = rule 26 |> Thm.lhs_of |> Thm.dest_binop |> pairself Thm.dest_arg;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   146
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   147
      val vars = [ca, cb, cc, cd, cm, cn, cp, cq, cx, cy, cz, clx, crx, cly, cry];
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   148
      val semiring = (sr_ops, sr_rules');
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   149
      val ring = (r_ops, r_rules');
30866
dd5117e2d73e now deals with devision in fields
chaieb
parents: 30722
diff changeset
   150
      val field = (f_ops, f_rules');
25254
0216ca99a599 Added field ideal into entry - uses by algebra method to prove the ideal membership problem
chaieb
parents: 24020
diff changeset
   151
      val ideal' = map (symmetric o mk_meta) ideal
23252
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   152
    in
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   153
      del_data key #>
25254
0216ca99a599 Added field ideal into entry - uses by algebra method to prove the ideal membership problem
chaieb
parents: 24020
diff changeset
   154
      apsnd (cons (key, ({vars = vars, semiring = semiring, 
30866
dd5117e2d73e now deals with devision in fields
chaieb
parents: 30722
diff changeset
   155
                          ring = ring, field = field, idom = idom, ideal = ideal'},
23335
42b827dfa86b Changed normalizer Data to contain a simpset besides the list of instances; Before the method starts, it simplifies with this simpset; Added attribute 'algebra' to add and delete theorems from the simpset above;
chaieb
parents: 23330
diff changeset
   156
             {is_const = undefined, dest_const = undefined, mk_const = undefined,
42b827dfa86b Changed normalizer Data to contain a simpset besides the list of instances; Before the method starts, it simplifies with this simpset; Added attribute 'algebra' to add and delete theorems from the simpset above;
chaieb
parents: 23330
diff changeset
   157
             conv = undefined})))
23252
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   158
    end);
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   159
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   160
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   161
(* extra-logical functions *)
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   162
23335
42b827dfa86b Changed normalizer Data to contain a simpset besides the list of instances; Before the method starts, it simplifies with this simpset; Added attribute 'algebra' to add and delete theorems from the simpset above;
chaieb
parents: 23330
diff changeset
   163
fun funs raw_key {is_const, dest_const, mk_const, conv} phi = 
42b827dfa86b Changed normalizer Data to contain a simpset besides the list of instances; Before the method starts, it simplifies with this simpset; Added attribute 'algebra' to add and delete theorems from the simpset above;
chaieb
parents: 23330
diff changeset
   164
 (Data.map o apsnd) (fn data =>
23252
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   165
  let
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   166
    val key = Morphism.thm phi raw_key;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   167
    val _ = AList.defined eq_key data key orelse
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   168
      raise THM ("No data entry for structure key", 0, [key]);
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   169
    val fns = {is_const = is_const phi, dest_const = dest_const phi,
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   170
      mk_const = mk_const phi, conv = conv phi};
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   171
  in AList.map_entry eq_key key (apsnd (K fns)) data end);
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   172
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   173
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   174
(* concrete syntax *)
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   175
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   176
local
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   177
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   178
fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K ();
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   179
fun keyword2 k1 k2 = Scan.lift (Args.$$$ k1 -- Args.$$$ k2 -- Args.colon) >> K ();
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   180
fun keyword3 k1 k2 k3 =
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   181
  Scan.lift (Args.$$$ k1 -- Args.$$$ k2 -- Args.$$$ k3 -- Args.colon) >> K ();
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   182
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   183
val opsN = "ops";
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   184
val rulesN = "rules";
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   185
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   186
val normN = "norm";
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   187
val constN = "const";
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   188
val delN = "del";
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   189
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   190
val any_keyword =
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   191
  keyword2 semiringN opsN || keyword2 semiringN rulesN ||
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   192
  keyword2 ringN opsN || keyword2 ringN rulesN ||
30866
dd5117e2d73e now deals with devision in fields
chaieb
parents: 30722
diff changeset
   193
  keyword2 fieldN opsN || keyword2 fieldN rulesN ||
25254
0216ca99a599 Added field ideal into entry - uses by algebra method to prove the ideal membership problem
chaieb
parents: 24020
diff changeset
   194
  keyword2 idomN rulesN || keyword2 idealN rulesN;
23252
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   195
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   196
val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   197
val terms = thms >> map Drule.dest_term;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   198
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   199
fun optional scan = Scan.optional scan [];
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   200
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   201
in
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   202
30722
623d4831c8cf simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents: 30528
diff changeset
   203
val normalizer_setup =
623d4831c8cf simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents: 30528
diff changeset
   204
  Attrib.setup @{binding normalizer}
623d4831c8cf simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents: 30528
diff changeset
   205
    (Scan.lift (Args.$$$ delN >> K del) ||
623d4831c8cf simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents: 30528
diff changeset
   206
      ((keyword2 semiringN opsN |-- terms) --
623d4831c8cf simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents: 30528
diff changeset
   207
       (keyword2 semiringN rulesN |-- thms)) --
623d4831c8cf simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents: 30528
diff changeset
   208
      (optional (keyword2 ringN opsN |-- terms) --
623d4831c8cf simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents: 30528
diff changeset
   209
       optional (keyword2 ringN rulesN |-- thms)) --
30866
dd5117e2d73e now deals with devision in fields
chaieb
parents: 30722
diff changeset
   210
      (optional (keyword2 fieldN opsN |-- terms) --
dd5117e2d73e now deals with devision in fields
chaieb
parents: 30722
diff changeset
   211
       optional (keyword2 fieldN rulesN |-- thms)) --
30722
623d4831c8cf simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents: 30528
diff changeset
   212
      optional (keyword2 idomN rulesN |-- thms) --
623d4831c8cf simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents: 30528
diff changeset
   213
      optional (keyword2 idealN rulesN |-- thms)
30866
dd5117e2d73e now deals with devision in fields
chaieb
parents: 30722
diff changeset
   214
      >> (fn ((((sr, r), f), id), idl) => 
dd5117e2d73e now deals with devision in fields
chaieb
parents: 30722
diff changeset
   215
             add {semiring = sr, ring = r, field = f, idom = id, ideal = idl}))
30722
623d4831c8cf simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents: 30528
diff changeset
   216
    "semiring normalizer data";
23252
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   217
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   218
end;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   219
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   220
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   221
(* theory setup *)
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   222
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   223
val setup =
30722
623d4831c8cf simplified attribute and method setup: eliminating bottom-up styles makes it easier to keep things in one place, and also SML/NJ happy;
wenzelm
parents: 30528
diff changeset
   224
  normalizer_setup #>
30528
7173bf123335 simplified attribute setup;
wenzelm
parents: 25979
diff changeset
   225
  Attrib.setup @{binding algebra} (Attrib.add_del add_ss del_ss) "pre-simplification for algebra";
23252
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   226
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   227
end;