| 24584 |      1 | (* Title:      HOL/Tools/Qelim/ferrante_rackoff_data.ML
 | 
| 23466 |      2 |    ID:         $Id$
 | 
|  |      3 |    Author:     Amine Chaieb, TU Muenchen
 | 
|  |      4 | 
 | 
|  |      5 | Context data for Ferrante and Rackoff's algorithm for quantifier
 | 
|  |      6 | elimination in dense linear orders.
 | 
|  |      7 | *)
 | 
|  |      8 | 
 | 
|  |      9 | signature FERRANTE_RACKOF_DATA =
 | 
|  |     10 | sig
 | 
|  |     11 |   datatype ord = Lt | Le | Gt | Ge | Eq | NEq | Nox
 | 
|  |     12 |   type entry
 | 
|  |     13 |   val get: Proof.context -> (thm * entry) list
 | 
|  |     14 |   val del: attribute
 | 
|  |     15 |   val add: entry -> attribute 
 | 
|  |     16 |   val funs: thm -> 
 | 
|  |     17 |     {isolate_conv: morphism -> Proof.context -> cterm list -> cterm -> thm,
 | 
|  |     18 |      whatis: morphism -> cterm -> cterm -> ord,
 | 
| 24020 |     19 |      simpset: morphism -> simpset} -> declaration
 | 
| 23466 |     20 |   val match: Proof.context -> cterm -> entry option
 | 
|  |     21 |   val setup: theory -> theory
 | 
|  |     22 | end;
 | 
|  |     23 | 
 | 
|  |     24 | structure Ferrante_Rackoff_Data: FERRANTE_RACKOF_DATA = 
 | 
|  |     25 | struct
 | 
|  |     26 | 
 | 
|  |     27 | (* data *)
 | 
|  |     28 | 
 | 
|  |     29 | datatype ord = Lt | Le | Gt | Ge | Eq | NEq | Nox
 | 
|  |     30 | 
 | 
|  |     31 | type entry = 
 | 
|  |     32 |   {minf: thm list, pinf: thm list, nmi: thm list, npi: thm list,  
 | 
|  |     33 |    ld: thm list, qe: thm, atoms : cterm list} *
 | 
|  |     34 |    {isolate_conv: Proof.context -> cterm list -> cterm -> thm, 
 | 
|  |     35 |     whatis : cterm -> cterm -> ord, 
 | 
|  |     36 |     simpset : simpset};
 | 
|  |     37 | 
 | 
|  |     38 | val eq_key = Thm.eq_thm;
 | 
|  |     39 | fun eq_data arg = eq_fst eq_key arg;
 | 
|  |     40 | 
 | 
|  |     41 | structure Data = GenericDataFun
 | 
|  |     42 | (
 | 
|  |     43 |   type T = (thm * entry) list;
 | 
|  |     44 |   val empty = [];
 | 
|  |     45 |   val extend = I;
 | 
|  |     46 |   fun merge _ = AList.merge eq_key (K true);
 | 
|  |     47 | );
 | 
|  |     48 | 
 | 
|  |     49 | val get = Data.get o Context.Proof;
 | 
|  |     50 | 
 | 
|  |     51 | fun del_data key = remove eq_data (key, []);
 | 
|  |     52 | 
 | 
|  |     53 | val del = Thm.declaration_attribute (Data.map o del_data);
 | 
|  |     54 | 
 | 
|  |     55 | fun undefined x = error "undefined";
 | 
|  |     56 | 
 | 
|  |     57 | fun add entry = 
 | 
|  |     58 |     Thm.declaration_attribute (fn key => fn context => context |> Data.map 
 | 
|  |     59 |       (del_data key #> cons (key, entry)));
 | 
|  |     60 | 
 | 
|  |     61 | 
 | 
|  |     62 | (* extra-logical functions *)
 | 
|  |     63 | 
 | 
|  |     64 | fun funs raw_key {isolate_conv = icv, whatis = wi, simpset = ss} phi = Data.map (fn data =>
 | 
|  |     65 |   let
 | 
|  |     66 |     val key = Morphism.thm phi raw_key;
 | 
|  |     67 |     val _ = AList.defined eq_key data key orelse
 | 
|  |     68 |       raise THM ("No data entry for structure key", 0, [key]);
 | 
|  |     69 |     val fns = {isolate_conv = icv phi, whatis = wi phi, simpset = ss phi};
 | 
|  |     70 |   in AList.map_entry eq_key key (apsnd (K fns)) data end);
 | 
|  |     71 | 
 | 
|  |     72 | fun match ctxt tm =
 | 
|  |     73 |   let
 | 
|  |     74 |     fun match_inst
 | 
|  |     75 |         ({minf, pinf, nmi, npi, ld, qe, atoms},
 | 
|  |     76 |          fns as {isolate_conv, whatis, simpset}) pat =
 | 
|  |     77 |        let
 | 
|  |     78 |         fun h instT =
 | 
|  |     79 |           let
 | 
|  |     80 |             val substT = Thm.instantiate (instT, []);
 | 
|  |     81 |             val substT_cterm = Drule.cterm_rule substT;
 | 
|  |     82 | 
 | 
|  |     83 |             val minf' = map substT minf
 | 
|  |     84 |             val pinf' = map substT pinf
 | 
|  |     85 |             val nmi' = map substT nmi
 | 
|  |     86 |             val npi' = map substT npi
 | 
|  |     87 |             val ld' = map substT ld
 | 
|  |     88 |             val qe' = substT qe
 | 
|  |     89 |             val atoms' = map substT_cterm atoms
 | 
|  |     90 |             val result = ({minf = minf', pinf = pinf', nmi = nmi', npi = npi', 
 | 
|  |     91 |                            ld = ld', qe = qe', atoms = atoms'}, fns)
 | 
|  |     92 |           in SOME result end
 | 
|  |     93 |       in (case try Thm.match (pat, tm) of
 | 
|  |     94 |            NONE => NONE
 | 
|  |     95 |          | SOME (instT, _) => h instT)
 | 
|  |     96 |       end;
 | 
|  |     97 | 
 | 
|  |     98 |     fun match_struct (_,
 | 
|  |     99 |         entry as ({atoms = atoms, ...}, _): entry) =
 | 
|  |    100 |       get_first (match_inst entry) atoms;
 | 
|  |    101 |   in get_first match_struct (get ctxt) end;
 | 
|  |    102 | 
 | 
|  |    103 | 
 | 
|  |    104 | (* concrete syntax *)
 | 
|  |    105 | 
 | 
|  |    106 | local
 | 
|  |    107 | val minfN = "minf";
 | 
|  |    108 | val pinfN = "pinf";
 | 
|  |    109 | val nmiN = "nmi";
 | 
|  |    110 | val npiN = "npi";
 | 
|  |    111 | val lin_denseN = "lindense";
 | 
|  |    112 | val qeN = "qe"
 | 
|  |    113 | val atomsN = "atoms"
 | 
|  |    114 | val simpsN = "simps"
 | 
|  |    115 | fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K ();
 | 
|  |    116 | val any_keyword =
 | 
|  |    117 |   keyword minfN || keyword pinfN || keyword nmiN 
 | 
|  |    118 | || keyword npiN || keyword lin_denseN || keyword qeN 
 | 
|  |    119 | || keyword atomsN || keyword simpsN;
 | 
|  |    120 | 
 | 
|  |    121 | val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat;
 | 
|  |    122 | val terms = thms >> map Drule.dest_term;
 | 
|  |    123 | in
 | 
|  |    124 | 
 | 
|  |    125 | fun att_syntax src = src |> Attrib.syntax
 | 
|  |    126 |     ((keyword minfN |-- thms)
 | 
|  |    127 |      -- (keyword pinfN |-- thms)
 | 
|  |    128 |      -- (keyword nmiN |-- thms)
 | 
|  |    129 |      -- (keyword npiN |-- thms)
 | 
|  |    130 |      -- (keyword lin_denseN |-- thms)
 | 
|  |    131 |      -- (keyword qeN |-- thms)
 | 
|  |    132 |      -- (keyword atomsN |-- terms) >> 
 | 
|  |    133 |      (fn ((((((minf,pinf),nmi),npi),lin_dense),qe), atoms)=> 
 | 
|  |    134 |      if length qe = 1 then
 | 
|  |    135 |        add ({minf = minf, pinf = pinf, nmi = nmi, npi = npi, ld = lin_dense, 
 | 
|  |    136 |             qe = hd qe, atoms = atoms},
 | 
|  |    137 |            {isolate_conv = undefined, whatis = undefined, simpset = HOL_ss})
 | 
|  |    138 |      else error "only one theorem for qe!"))
 | 
|  |    139 | 
 | 
|  |    140 | end;
 | 
|  |    141 | 
 | 
|  |    142 | 
 | 
|  |    143 | (* theory setup *)
 | 
|  |    144 | 
 | 
|  |    145 | val setup =
 | 
| 23903 |    146 |   Attrib.add_attributes [("ferrack", att_syntax, "Ferrante Rackoff data")];
 | 
| 23466 |    147 | 
 | 
|  |    148 | end;
 |