src/HOL/Tools/Groebner_Basis/normalizer_data.ML
author haftmann
Mon, 16 Jul 2007 09:29:04 +0200
changeset 23811 b18557301bf9
parent 23581 297c6d706322
child 24020 ed4d7abffee7
permissions -rw-r--r--
added function for case certificates
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
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    14
  val add: {semiring: cterm list * thm list, ring: cterm list * thm list, idom: thm list}
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,
23485
881b04972953 made type conv pervasive;
wenzelm
parents: 23335
diff changeset
    19
    conv: morphism -> Proof.context -> cterm -> thm} ->
881b04972953 made type conv pervasive;
wenzelm
parents: 23335
diff changeset
    20
      morphism -> Context.generic -> Context.generic
23252
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    21
  val setup: theory -> theory
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    22
end;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    23
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    24
structure NormalizerData: NORMALIZER_DATA =
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    25
struct
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    26
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    27
(* data *)
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    28
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    29
type entry =
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    30
 {vars: cterm list,
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    31
  semiring: cterm list * thm list,
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    32
  ring: cterm list * thm list,
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    33
  idom: thm list} *
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    34
 {is_const: cterm -> bool,
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    35
  dest_const: cterm -> Rat.rat,
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    36
  mk_const: ctyp -> Rat.rat -> cterm,
23330
01c09922ce59 Conversion for computation on constants now depends on the context
chaieb
parents: 23260
diff changeset
    37
  conv: Proof.context -> cterm -> thm};
23252
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    38
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    39
val eq_key = Thm.eq_thm;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    40
fun eq_data arg = eq_fst eq_key arg;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    41
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    42
structure Data = GenericDataFun
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    43
(
23581
wenzelm
parents: 23485
diff changeset
    44
  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
    45
  val empty = (HOL_basic_ss, []);
23252
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    46
  val extend = I;
23581
wenzelm
parents: 23485
diff changeset
    47
  fun merge _ ((ss, e), (ss', e')) =
wenzelm
parents: 23485
diff changeset
    48
    (merge_ss (ss, ss'), AList.merge eq_key (K true) (e, e'));
23252
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    49
);
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    50
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    51
val get = Data.get o Context.Proof;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    52
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    53
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    54
(* match data *)
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    55
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    56
fun match ctxt tm =
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    57
  let
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    58
    fun match_inst
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    59
        ({vars, semiring = (sr_ops, sr_rules), ring = (r_ops, r_rules), idom},
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    60
         fns as {is_const, dest_const, mk_const, conv}) pat =
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    61
       let
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    62
        fun h instT =
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    63
          let
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    64
            val substT = Thm.instantiate (instT, []);
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    65
            val substT_cterm = Drule.cterm_rule substT;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    66
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    67
            val vars' = map substT_cterm vars;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    68
            val semiring' = (map substT_cterm sr_ops, map substT sr_rules);
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    69
            val ring' = (map substT_cterm r_ops, map substT r_rules);
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    70
            val idom' = map substT idom;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    71
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    72
            val result = ({vars = vars', semiring = semiring', ring = ring', idom = idom'}, fns);
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    73
          in SOME result end
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    74
      in (case try Thm.match (pat, tm) of
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    75
           NONE => NONE
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    76
         | SOME (instT, _) => h instT)
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    77
      end;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    78
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    79
    fun match_struct (_,
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    80
        entry as ({semiring = (sr_ops, _), ring = (r_ops, _), ...}, _): entry) =
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    81
      get_first (match_inst entry) (sr_ops @ r_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
    82
  in get_first match_struct (snd (get ctxt)) end;
23252
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    83
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    84
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    85
(* logical content *)
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    86
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    87
val semiringN = "semiring";
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    88
val ringN = "ring";
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    89
val idomN = "idom";
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    90
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    91
fun undefined _ = raise Match;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    92
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
    93
fun del_data key = apsnd (remove eq_data (key, []));
23252
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    94
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
    95
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
    96
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
    97
   (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
    98
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
    99
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
   100
   (fn th => Data.map (fn (ss,data) => (ss delsimps [th], data)));
23252
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   101
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   102
fun add {semiring = (sr_ops, sr_rules), ring = (r_ops, r_rules), idom} =
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   103
  Thm.declaration_attribute (fn key => fn context => context |> Data.map
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   104
    let
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   105
      val ctxt = Context.proof_of context;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   106
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   107
      fun check kind name xs n =
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   108
        null xs orelse length xs = n orelse
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   109
        error ("Expected " ^ string_of_int n ^ " " ^ kind ^ " for " ^ name);
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   110
      val check_ops = check "operations";
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   111
      val check_rules = check "rules";
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   112
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   113
      val _ =
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   114
        check_ops semiringN sr_ops 5 andalso
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   115
        check_rules semiringN sr_rules 37 andalso
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   116
        check_ops ringN r_ops 2 andalso
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   117
        check_rules ringN r_rules 2 andalso
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   118
        check_rules idomN idom 2;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   119
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   120
      val mk_meta = LocalDefs.meta_rewrite_rule ctxt;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   121
      val sr_rules' = map mk_meta sr_rules;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   122
      val r_rules' = map mk_meta r_rules;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   123
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   124
      fun rule i = nth sr_rules' (i - 1);
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   125
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   126
      val (cx, cy) = Thm.dest_binop (hd sr_ops);
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   127
      val cz = rule 34 |> Thm.rhs_of |> Thm.dest_arg |> Thm.dest_arg;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   128
      val cn = rule 36 |> Thm.rhs_of |> Thm.dest_arg |> Thm.dest_arg;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   129
      val ((clx, crx), (cly, cry)) =
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   130
        rule 13 |> Thm.rhs_of |> Thm.dest_binop |> pairself Thm.dest_binop;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   131
      val ((ca, cb), (cc, cd)) =
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   132
        rule 20 |> Thm.lhs_of |> Thm.dest_binop |> pairself Thm.dest_binop;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   133
      val cm = rule 1 |> Thm.rhs_of |> Thm.dest_arg;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   134
      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
   135
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   136
      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
   137
      val semiring = (sr_ops, sr_rules');
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   138
      val ring = (r_ops, r_rules');
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   139
    in
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   140
      del_data key #>
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
   141
      apsnd (cons (key, ({vars = vars, semiring = semiring, ring = ring, idom = idom},
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
   142
             {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
   143
             conv = undefined})))
23252
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   144
    end);
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   145
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   146
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   147
(* extra-logical functions *)
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   148
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
   149
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
   150
 (Data.map o apsnd) (fn data =>
23252
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   151
  let
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   152
    val key = Morphism.thm phi raw_key;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   153
    val _ = AList.defined eq_key data key orelse
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   154
      raise THM ("No data entry for structure key", 0, [key]);
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   155
    val fns = {is_const = is_const phi, dest_const = dest_const phi,
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   156
      mk_const = mk_const phi, conv = conv phi};
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   157
  in AList.map_entry eq_key key (apsnd (K fns)) data end);
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   158
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   159
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   160
(* concrete syntax *)
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   161
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   162
local
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   163
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   164
fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K ();
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   165
fun keyword2 k1 k2 = Scan.lift (Args.$$$ k1 -- Args.$$$ k2 -- Args.colon) >> K ();
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   166
fun keyword3 k1 k2 k3 =
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   167
  Scan.lift (Args.$$$ k1 -- Args.$$$ k2 -- Args.$$$ k3 -- Args.colon) >> K ();
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   168
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   169
val opsN = "ops";
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   170
val rulesN = "rules";
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   171
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   172
val normN = "norm";
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   173
val constN = "const";
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   174
val delN = "del";
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   175
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   176
val any_keyword =
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   177
  keyword2 semiringN opsN || keyword2 semiringN rulesN ||
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   178
  keyword2 ringN opsN || keyword2 ringN rulesN ||
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   179
  keyword2 idomN rulesN;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   180
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   181
val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   182
val terms = thms >> map Drule.dest_term;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   183
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   184
fun optional scan = Scan.optional scan [];
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   185
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   186
in
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   187
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   188
fun att_syntax src = src |> Attrib.syntax
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   189
  (Scan.lift (Args.$$$ delN >> K del) ||
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   190
    ((keyword2 semiringN opsN |-- terms) --
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   191
     (keyword2 semiringN rulesN |-- thms)) --
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   192
    (optional (keyword2 ringN opsN |-- terms) --
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   193
     optional (keyword2 ringN rulesN |-- thms)) --
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   194
    optional (keyword2 idomN rulesN |-- thms)
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   195
    >> (fn ((sr, r), id) => add {semiring = sr, ring = r, idom = id}));
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   196
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   197
end;
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   198
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   199
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   200
(* theory setup *)
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   201
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   202
val setup =
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
   203
  Attrib.add_attributes 
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
   204
     [("normalizer", att_syntax, "semiring normalizer 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
   205
      ("algebra", Attrib.add_del_args add_ss del_ss, "Pre-simplification for algebra")];
23252
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   206
67268bb40b21 Semiring normalization and Groebner Bases.
wenzelm
parents:
diff changeset
   207
end;