| author | wenzelm | 
| Fri, 04 Jun 2010 11:30:46 +0200 | |
| changeset 37311 | 90323e435a7f | 
| parent 37120 | 5df087c6ce94 | 
| child 41226 | adcb9a1198e7 | 
| permissions | -rw-r--r-- | 
| 23907 | 1  | 
signature LANGFORD_DATA =  | 
2  | 
sig  | 
|
3  | 
type entry  | 
|
4  | 
val get: Proof.context -> simpset * (thm * entry) list  | 
|
5  | 
val add: entry -> attribute  | 
|
6  | 
val del: attribute  | 
|
7  | 
val setup: theory -> theory  | 
|
8  | 
val match: Proof.context -> cterm -> entry option  | 
|
9  | 
end;  | 
|
10  | 
||
11  | 
structure Langford_Data: LANGFORD_DATA =  | 
|
12  | 
struct  | 
|
13  | 
||
14  | 
(* data *)  | 
|
15  | 
type entry = {qe_bnds: thm, qe_nolb : thm , qe_noub: thm, 
 | 
|
16  | 
gs : thm list, gst : thm, atoms: cterm list};  | 
|
17  | 
val eq_key = Thm.eq_thm;  | 
|
18  | 
fun eq_data arg = eq_fst eq_key arg;  | 
|
19  | 
||
| 33519 | 20  | 
structure Data = Generic_Data  | 
21  | 
(  | 
|
22  | 
type T = simpset * (thm * entry) list;  | 
|
| 23907 | 23  | 
val empty = (HOL_basic_ss, []);  | 
24  | 
val extend = I;  | 
|
| 33519 | 25  | 
fun merge ((ss1, es1), (ss2, es2)) : T =  | 
26  | 
(merge_ss (ss1, ss2), AList.merge eq_key (K true) (es1, es2));  | 
|
| 23907 | 27  | 
);  | 
28  | 
||
29  | 
val get = Data.get o Context.Proof;  | 
|
30  | 
||
31  | 
fun del_data key = apsnd (remove eq_data (key, []));  | 
|
32  | 
||
33  | 
val del = Thm.declaration_attribute (Data.map o del_data);  | 
|
34  | 
||
35  | 
fun add entry =  | 
|
36  | 
Thm.declaration_attribute (fn key => fn context => context |> Data.map  | 
|
37  | 
(del_data key #> apsnd (cons (key, entry))));  | 
|
38  | 
||
39  | 
val add_simp = Thm.declaration_attribute (fn th => fn context =>  | 
|
40  | 
context |> Data.map (fn (ss,ts') => (ss addsimps [th], ts')))  | 
|
41  | 
||
42  | 
val del_simp = Thm.declaration_attribute (fn th => fn context =>  | 
|
43  | 
context |> Data.map (fn (ss,ts') => (ss delsimps [th], ts')))  | 
|
44  | 
||
45  | 
fun match ctxt tm =  | 
|
46  | 
let  | 
|
47  | 
fun match_inst  | 
|
48  | 
        {qe_bnds, qe_nolb, qe_noub, gs, gst, atoms} pat =
 | 
|
49  | 
let  | 
|
50  | 
fun h instT =  | 
|
51  | 
let  | 
|
52  | 
val substT = Thm.instantiate (instT, []);  | 
|
53  | 
val substT_cterm = Drule.cterm_rule substT;  | 
|
54  | 
||
55  | 
val qe_bnds' = substT qe_bnds  | 
|
56  | 
val qe_nolb' = substT qe_nolb  | 
|
57  | 
val qe_noub' = substT qe_noub  | 
|
58  | 
val gs' = map substT gs  | 
|
59  | 
val gst' = substT gst  | 
|
60  | 
val atoms' = map substT_cterm atoms  | 
|
61  | 
            val result = {qe_bnds = qe_bnds', qe_nolb = qe_nolb', 
 | 
|
62  | 
qe_noub = qe_noub', gs = gs', gst = gst',  | 
|
63  | 
atoms = atoms'}  | 
|
64  | 
in SOME result end  | 
|
65  | 
in (case try Thm.match (pat, tm) of  | 
|
66  | 
NONE => NONE  | 
|
67  | 
| SOME (instT, _) => h instT)  | 
|
68  | 
end;  | 
|
69  | 
||
70  | 
fun match_struct (_,  | 
|
71  | 
        entry as ({atoms = atoms, ...}): entry) =
 | 
|
72  | 
get_first (match_inst entry) atoms;  | 
|
73  | 
in get_first match_struct (snd (get ctxt)) end;  | 
|
74  | 
||
75  | 
(* concrete syntax *)  | 
|
76  | 
||
77  | 
local  | 
|
78  | 
val qeN = "qe";  | 
|
79  | 
val gatherN = "gather";  | 
|
80  | 
val atomsN = "atoms"  | 
|
81  | 
fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K ();  | 
|
82  | 
val any_keyword =  | 
|
83  | 
keyword qeN || keyword gatherN || keyword atomsN;  | 
|
84  | 
||
85  | 
val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat;  | 
|
86  | 
val terms = thms >> map Drule.dest_term;  | 
|
87  | 
in  | 
|
88  | 
||
| 
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
 | 
89  | 
val langford_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
 | 
90  | 
  Attrib.setup @{binding langford}
 | 
| 
 
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
 | 
91  | 
((keyword qeN |-- thms)  | 
| 23907 | 92  | 
-- (keyword gatherN |-- 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
 | 
93  | 
-- (keyword atomsN |-- 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
 | 
94  | 
(fn ((qes,gs), atoms) =>  | 
| 
 
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
 | 
95  | 
if length qes = 3 andalso length gs > 1 then  | 
| 
 
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
 | 
96  | 
let  | 
| 
 
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
 | 
97  | 
val [q1,q2,q3] = qes  | 
| 
 
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
 | 
98  | 
val gst::gss = gs  | 
| 
 
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
 | 
99  | 
           val entry = {qe_bnds = q1, qe_nolb = q2,
 | 
| 
 
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
 | 
100  | 
qe_noub = q3, gs = gss, gst = gst, atoms = atoms}  | 
| 
 
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
 | 
101  | 
in add entry end  | 
| 
 
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
 | 
102  | 
else error "Need 3 theorems for qe and at least one for gs"))  | 
| 
 
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
 | 
103  | 
"Langford 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: 
30528 
diff
changeset
 | 
104  | 
|
| 23907 | 105  | 
end;  | 
106  | 
||
107  | 
(* theory setup *)  | 
|
108  | 
||
109  | 
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
 | 
110  | 
langford_setup #>  | 
| 30528 | 111  | 
  Attrib.setup @{binding langfordsimp} (Attrib.add_del add_simp del_simp) "Langford simpset";
 | 
| 23907 | 112  | 
|
113  | 
end;  |