src/HOL/Tools/Qelim/cooper_data.ML
author wenzelm
Wed, 31 Dec 2008 18:53:16 +0100
changeset 29270 0eade173f77e
parent 26086 3c243098b64a
child 30528 7173bf123335
permissions -rw-r--r--
moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
24584
01e83ffa6c54 fixed title
haftmann
parents: 23879
diff changeset
     1
(*  Title:      HOL/Tools/Qelim/cooper_data.ML
23466
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
     3
    Author:     Amine Chaieb, TU Muenchen
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
     4
*)
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
     5
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
     6
signature COOPER_DATA =
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
     7
sig
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
     8
  type entry
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
     9
  val get: Proof.context -> entry
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    10
  val del: term list -> attribute
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    11
  val add: term list -> attribute 
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    12
  val setup: theory -> theory
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    13
end;
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    14
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    15
structure CooperData : COOPER_DATA =
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    16
struct
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    17
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    18
type entry = simpset * (term list);
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    19
val start_ss = HOL_ss (* addsimps @{thms "Groebner_Basis.comp_arith"}
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    20
               addcongs [if_weak_cong, @{thm "let_weak_cong"}];*)
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    21
val allowed_consts = 
25762
c03e9d04b3e4 splitted class uminus from class minus
haftmann
parents: 24584
diff changeset
    22
  [@{term "op + :: int => _"}, @{term "op + :: nat => _"},
c03e9d04b3e4 splitted class uminus from class minus
haftmann
parents: 24584
diff changeset
    23
   @{term "op - :: int => _"}, @{term "op - :: nat => _"},
c03e9d04b3e4 splitted class uminus from class minus
haftmann
parents: 24584
diff changeset
    24
   @{term "op * :: int => _"}, @{term "op * :: nat => _"},
c03e9d04b3e4 splitted class uminus from class minus
haftmann
parents: 24584
diff changeset
    25
   @{term "op div :: int => _"}, @{term "op div :: nat => _"},
c03e9d04b3e4 splitted class uminus from class minus
haftmann
parents: 24584
diff changeset
    26
   @{term "op mod :: int => _"}, @{term "op mod :: nat => _"},
26086
3c243098b64a New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents: 25979
diff changeset
    27
   @{term "Int.Bit0"}, @{term "Int.Bit1"},
23466
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    28
   @{term "op &"}, @{term "op |"}, @{term "op -->"}, 
25762
c03e9d04b3e4 splitted class uminus from class minus
haftmann
parents: 24584
diff changeset
    29
   @{term "op = :: int => _"}, @{term "op = :: nat => _"}, @{term "op = :: bool => _"},
23466
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    30
   @{term "op < :: int => _"}, @{term "op < :: nat => _"},
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    31
   @{term "op <= :: int => _"}, @{term "op <= :: nat => _"},
25762
c03e9d04b3e4 splitted class uminus from class minus
haftmann
parents: 24584
diff changeset
    32
   @{term "op dvd :: int => _"}, @{term "op dvd :: nat => _"},
c03e9d04b3e4 splitted class uminus from class minus
haftmann
parents: 24584
diff changeset
    33
   @{term "abs :: int => _"},
c03e9d04b3e4 splitted class uminus from class minus
haftmann
parents: 24584
diff changeset
    34
   @{term "max :: int => _"}, @{term "max :: nat => _"},
c03e9d04b3e4 splitted class uminus from class minus
haftmann
parents: 24584
diff changeset
    35
   @{term "min :: int => _"}, @{term "min :: nat => _"},
c03e9d04b3e4 splitted class uminus from class minus
haftmann
parents: 24584
diff changeset
    36
   @{term "HOL.uminus :: int => _"}, (*@ {term "HOL.uminus :: nat => _"},*)
c03e9d04b3e4 splitted class uminus from class minus
haftmann
parents: 24584
diff changeset
    37
   @{term "Not"}, @{term "Suc"},
c03e9d04b3e4 splitted class uminus from class minus
haftmann
parents: 24584
diff changeset
    38
   @{term "Ex :: (int => _) => _"}, @{term "Ex :: (nat => _) => _"},
c03e9d04b3e4 splitted class uminus from class minus
haftmann
parents: 24584
diff changeset
    39
   @{term "All :: (int => _) => _"}, @{term "All :: (nat => _) => _"},
23466
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    40
   @{term "nat"}, @{term "int"},
26086
3c243098b64a New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents: 25979
diff changeset
    41
   @{term "Int.Bit0"}, @{term "Int.Bit1"},
3c243098b64a New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents: 25979
diff changeset
    42
   @{term "Int.Pls"}, @{term "Int.Min"},
25919
8b1c0d434824 joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents: 25762
diff changeset
    43
   @{term "Int.number_of :: int => int"}, @{term "Int.number_of :: int => nat"},
23466
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    44
   @{term "0::int"}, @{term "1::int"}, @{term "0::nat"}, @{term "1::nat"},
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    45
   @{term "True"}, @{term "False"}];
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    46
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    47
structure Data = GenericDataFun
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    48
(
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    49
  type T = simpset * (term list);
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    50
  val empty = (start_ss, allowed_consts);
23474
688987d0bab4 merge is now identity
chaieb
parents: 23466
diff changeset
    51
  val extend  = I;
23466
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    52
  fun merge _ ((ss1, ts1), (ss2, ts2)) =
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    53
    (merge_ss (ss1, ss2), Library.merge (op aconv) (ts1, ts2));
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    54
);
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    55
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    56
val get = Data.get o Context.Proof;
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    57
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    58
fun add ts = Thm.declaration_attribute (fn th => fn context => 
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    59
  context |> Data.map (fn (ss,ts') => 
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    60
     (ss addsimps [th], merge (op aconv) (ts',ts) ))) 
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    61
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    62
fun del ts = Thm.declaration_attribute (fn th => fn context => 
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    63
  context |> Data.map (fn (ss,ts') => 
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    64
     (ss delsimps [th], subtract (op aconv) ts' ts ))) 
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    65
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    66
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    67
(* concrete syntax *)
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    68
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    69
local
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    70
fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K ();
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    71
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    72
val constsN = "consts";
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    73
val any_keyword = keyword constsN
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    74
val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat;
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    75
val terms = thms >> map (term_of o Drule.dest_term);
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    76
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    77
fun optional scan = Scan.optional scan [];
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    78
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    79
in
25979
3297781f8141 tuned attribute syntax -- no need for eta-expansion;
wenzelm
parents: 25919
diff changeset
    80
val att_syntax = Attrib.syntax
23466
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    81
 ((Scan.lift (Args.$$$ "del") |-- optional (keyword constsN |-- terms)) >> del ||
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    82
  optional (keyword constsN |-- terms) >> add)
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    83
end;
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    84
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    85
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    86
(* theory setup *)
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    87
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    88
val setup =
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    89
  Attrib.add_attributes [("presburger", att_syntax, "Cooper data")];
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    90
886655a150f6 moved quantifier elimination tools to Tools/Qelim/;
wenzelm
parents:
diff changeset
    91
end;