| author | haftmann | 
| Fri, 24 Apr 2009 18:20:37 +0200 | |
| changeset 30975 | b2fa60d56735 | 
| parent 30722 | 623d4831c8cf | 
| child 33519 | e31a85f92ce9 | 
| permissions | -rw-r--r-- | 
| 24584 | 1 | (* Title: HOL/Tools/Qelim/cooper_data.ML | 
| 23466 | 2 | ID: $Id$ | 
| 3 | Author: Amine Chaieb, TU Muenchen | |
| 4 | *) | |
| 5 | ||
| 6 | signature COOPER_DATA = | |
| 7 | sig | |
| 8 | type entry | |
| 9 | val get: Proof.context -> entry | |
| 10 | val del: term list -> attribute | |
| 11 | val add: term list -> attribute | |
| 12 | val setup: theory -> theory | |
| 13 | end; | |
| 14 | ||
| 15 | structure CooperData : COOPER_DATA = | |
| 16 | struct | |
| 17 | ||
| 18 | type entry = simpset * (term list); | |
| 19 | val start_ss = HOL_ss (* addsimps @{thms "Groebner_Basis.comp_arith"}
 | |
| 20 |                addcongs [if_weak_cong, @{thm "let_weak_cong"}];*)
 | |
| 21 | val allowed_consts = | |
| 25762 | 22 |   [@{term "op + :: int => _"}, @{term "op + :: nat => _"},
 | 
| 23 |    @{term "op - :: int => _"}, @{term "op - :: nat => _"},
 | |
| 24 |    @{term "op * :: int => _"}, @{term "op * :: nat => _"},
 | |
| 25 |    @{term "op div :: int => _"}, @{term "op div :: nat => _"},
 | |
| 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: 
25979diff
changeset | 27 |    @{term "Int.Bit0"}, @{term "Int.Bit1"},
 | 
| 23466 | 28 |    @{term "op &"}, @{term "op |"}, @{term "op -->"}, 
 | 
| 25762 | 29 |    @{term "op = :: int => _"}, @{term "op = :: nat => _"}, @{term "op = :: bool => _"},
 | 
| 23466 | 30 |    @{term "op < :: int => _"}, @{term "op < :: nat => _"},
 | 
| 31 |    @{term "op <= :: int => _"}, @{term "op <= :: nat => _"},
 | |
| 25762 | 32 |    @{term "op dvd :: int => _"}, @{term "op dvd :: nat => _"},
 | 
| 33 |    @{term "abs :: int => _"},
 | |
| 34 |    @{term "max :: int => _"}, @{term "max :: nat => _"},
 | |
| 35 |    @{term "min :: int => _"}, @{term "min :: nat => _"},
 | |
| 36 |    @{term "HOL.uminus :: int => _"}, (*@ {term "HOL.uminus :: nat => _"},*)
 | |
| 37 |    @{term "Not"}, @{term "Suc"},
 | |
| 38 |    @{term "Ex :: (int => _) => _"}, @{term "Ex :: (nat => _) => _"},
 | |
| 39 |    @{term "All :: (int => _) => _"}, @{term "All :: (nat => _) => _"},
 | |
| 23466 | 40 |    @{term "nat"}, @{term "int"},
 | 
| 26086 
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
 huffman parents: 
25979diff
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: 
25979diff
changeset | 42 |    @{term "Int.Pls"}, @{term "Int.Min"},
 | 
| 25919 
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
 haftmann parents: 
25762diff
changeset | 43 |    @{term "Int.number_of :: int => int"}, @{term "Int.number_of :: int => nat"},
 | 
| 23466 | 44 |    @{term "0::int"}, @{term "1::int"}, @{term "0::nat"}, @{term "1::nat"},
 | 
| 45 |    @{term "True"}, @{term "False"}];
 | |
| 46 | ||
| 47 | structure Data = GenericDataFun | |
| 48 | ( | |
| 49 | type T = simpset * (term list); | |
| 50 | val empty = (start_ss, allowed_consts); | |
| 23474 | 51 | val extend = I; | 
| 23466 | 52 | fun merge _ ((ss1, ts1), (ss2, ts2)) = | 
| 53 | (merge_ss (ss1, ss2), Library.merge (op aconv) (ts1, ts2)); | |
| 54 | ); | |
| 55 | ||
| 56 | val get = Data.get o Context.Proof; | |
| 57 | ||
| 58 | fun add ts = Thm.declaration_attribute (fn th => fn context => | |
| 59 | context |> Data.map (fn (ss,ts') => | |
| 60 | (ss addsimps [th], merge (op aconv) (ts',ts) ))) | |
| 61 | ||
| 62 | fun del ts = Thm.declaration_attribute (fn th => fn context => | |
| 63 | context |> Data.map (fn (ss,ts') => | |
| 64 | (ss delsimps [th], subtract (op aconv) ts' ts ))) | |
| 65 | ||
| 66 | ||
| 67 | (* concrete syntax *) | |
| 68 | ||
| 69 | local | |
| 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: 
30528diff
changeset | 70 | |
| 23466 | 71 | fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K (); | 
| 72 | ||
| 73 | val constsN = "consts"; | |
| 74 | val any_keyword = keyword constsN | |
| 75 | val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat; | |
| 76 | val terms = thms >> map (term_of o Drule.dest_term); | |
| 77 | ||
| 78 | fun optional scan = Scan.optional scan []; | |
| 79 | ||
| 80 | in | |
| 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: 
30528diff
changeset | 81 | |
| 
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: 
30528diff
changeset | 82 | val presburger_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: 
30528diff
changeset | 83 |   Attrib.setup @{binding presburger}
 | 
| 
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: 
30528diff
changeset | 84 | ((Scan.lift (Args.$$$ "del") |-- optional (keyword constsN |-- terms)) >> 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: 
30528diff
changeset | 85 | optional (keyword constsN |-- terms) >> add) "Cooper data"; | 
| 
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: 
30528diff
changeset | 86 | |
| 23466 | 87 | end; | 
| 88 | ||
| 89 | ||
| 90 | (* theory setup *) | |
| 91 | ||
| 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: 
30528diff
changeset | 92 | val setup = presburger_setup; | 
| 23466 | 93 | |
| 94 | end; |