author | haftmann |
Sat, 03 Mar 2012 22:38:33 +0100 | |
changeset 46790 | f3c10e908f65 |
parent 41226 | adcb9a1198e7 |
child 51717 | 9e7d1c139569 |
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 |
||
41226 | 39 |
val add_simp = Thm.declaration_attribute (Data.map o apfst o Simplifier.add_simp); |
23907 | 40 |
|
41226 | 41 |
val del_simp = Thm.declaration_attribute (Data.map o apfst o Simplifier.del_simp); |
23907 | 42 |
|
43 |
fun match ctxt tm = |
|
44 |
let |
|
45 |
fun match_inst |
|
46 |
{qe_bnds, qe_nolb, qe_noub, gs, gst, atoms} pat = |
|
47 |
let |
|
48 |
fun h instT = |
|
49 |
let |
|
50 |
val substT = Thm.instantiate (instT, []); |
|
51 |
val substT_cterm = Drule.cterm_rule substT; |
|
52 |
||
53 |
val qe_bnds' = substT qe_bnds |
|
54 |
val qe_nolb' = substT qe_nolb |
|
55 |
val qe_noub' = substT qe_noub |
|
56 |
val gs' = map substT gs |
|
57 |
val gst' = substT gst |
|
58 |
val atoms' = map substT_cterm atoms |
|
59 |
val result = {qe_bnds = qe_bnds', qe_nolb = qe_nolb', |
|
60 |
qe_noub = qe_noub', gs = gs', gst = gst', |
|
61 |
atoms = atoms'} |
|
62 |
in SOME result end |
|
63 |
in (case try Thm.match (pat, tm) of |
|
64 |
NONE => NONE |
|
65 |
| SOME (instT, _) => h instT) |
|
66 |
end; |
|
67 |
||
68 |
fun match_struct (_, |
|
69 |
entry as ({atoms = atoms, ...}): entry) = |
|
70 |
get_first (match_inst entry) atoms; |
|
71 |
in get_first match_struct (snd (get ctxt)) end; |
|
72 |
||
73 |
(* concrete syntax *) |
|
74 |
||
75 |
local |
|
76 |
val qeN = "qe"; |
|
77 |
val gatherN = "gather"; |
|
78 |
val atomsN = "atoms" |
|
79 |
fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K (); |
|
80 |
val any_keyword = |
|
81 |
keyword qeN || keyword gatherN || keyword atomsN; |
|
82 |
||
83 |
val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat; |
|
84 |
val terms = thms >> map Drule.dest_term; |
|
85 |
in |
|
86 |
||
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
|
87 |
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
|
88 |
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
|
89 |
((keyword qeN |-- thms) |
23907 | 90 |
-- (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
|
91 |
-- (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
|
92 |
(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
|
93 |
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
|
94 |
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
|
95 |
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
|
96 |
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
|
97 |
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
|
98 |
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
|
99 |
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
|
100 |
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
|
101 |
"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
|
102 |
|
23907 | 103 |
end; |
104 |
||
105 |
(* theory setup *) |
|
106 |
||
107 |
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
|
108 |
langford_setup #> |
30528 | 109 |
Attrib.setup @{binding langfordsimp} (Attrib.add_del add_simp del_simp) "Langford simpset"; |
23907 | 110 |
|
111 |
end; |