| author | haftmann | 
| Thu, 23 Apr 2009 12:17:51 +0200 | |
| changeset 30968 | 10fef94f40fc | 
| parent 30722 | 623d4831c8cf | 
| child 33519 | e31a85f92ce9 | 
| 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 | ||
| 20 | structure Data = GenericDataFun | |
| 21 | ( type T = simpset * (thm * entry) list; | |
| 22 | val empty = (HOL_basic_ss, []); | |
| 23 | val extend = I; | |
| 24 | fun merge _ ((ss1,es1), (ss2,es2)) = | |
| 25 | (merge_ss (ss1,ss2), AList.merge eq_key (K true) (es1,es2));; | |
| 26 | ); | |
| 27 | ||
| 28 | val get = Data.get o Context.Proof; | |
| 29 | ||
| 30 | fun del_data key = apsnd (remove eq_data (key, [])); | |
| 31 | ||
| 32 | val del = Thm.declaration_attribute (Data.map o del_data); | |
| 33 | ||
| 34 | fun add entry = | |
| 35 | Thm.declaration_attribute (fn key => fn context => context |> Data.map | |
| 36 | (del_data key #> apsnd (cons (key, entry)))); | |
| 37 | ||
| 38 | val add_simp = Thm.declaration_attribute (fn th => fn context => | |
| 39 | context |> Data.map (fn (ss,ts') => (ss addsimps [th], ts'))) | |
| 40 | ||
| 41 | val del_simp = Thm.declaration_attribute (fn th => fn context => | |
| 42 | context |> Data.map (fn (ss,ts') => (ss delsimps [th], ts'))) | |
| 43 | ||
| 44 | fun match ctxt tm = | |
| 45 | let | |
| 46 | fun match_inst | |
| 47 |         {qe_bnds, qe_nolb, qe_noub, gs, gst, atoms} pat =
 | |
| 48 | let | |
| 49 | fun h instT = | |
| 50 | let | |
| 51 | val substT = Thm.instantiate (instT, []); | |
| 52 | val substT_cterm = Drule.cterm_rule substT; | |
| 53 | ||
| 54 | val qe_bnds' = substT qe_bnds | |
| 55 | val qe_nolb' = substT qe_nolb | |
| 56 | val qe_noub' = substT qe_noub | |
| 57 | val gs' = map substT gs | |
| 58 | val gst' = substT gst | |
| 59 | val atoms' = map substT_cterm atoms | |
| 60 |             val result = {qe_bnds = qe_bnds', qe_nolb = qe_nolb', 
 | |
| 61 | qe_noub = qe_noub', gs = gs', gst = gst', | |
| 62 | atoms = atoms'} | |
| 63 | in SOME result end | |
| 64 | in (case try Thm.match (pat, tm) of | |
| 65 | NONE => NONE | |
| 66 | | SOME (instT, _) => h instT) | |
| 67 | end; | |
| 68 | ||
| 69 | fun match_struct (_, | |
| 70 |         entry as ({atoms = atoms, ...}): entry) =
 | |
| 71 | get_first (match_inst entry) atoms; | |
| 72 | in get_first match_struct (snd (get ctxt)) end; | |
| 73 | ||
| 74 | (* concrete syntax *) | |
| 75 | ||
| 76 | local | |
| 77 | val qeN = "qe"; | |
| 78 | val gatherN = "gather"; | |
| 79 | val atomsN = "atoms" | |
| 80 | fun keyword k = Scan.lift (Args.$$$ k -- Args.colon) >> K (); | |
| 81 | val any_keyword = | |
| 82 | keyword qeN || keyword gatherN || keyword atomsN; | |
| 83 | ||
| 84 | val thms = Scan.repeat (Scan.unless any_keyword Attrib.multi_thm) >> flat; | |
| 85 | val terms = thms >> map Drule.dest_term; | |
| 86 | in | |
| 87 | ||
| 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 | 88 | 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: 
30528diff
changeset | 89 |   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: 
30528diff
changeset | 90 | ((keyword qeN |-- thms) | 
| 23907 | 91 | -- (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: 
30528diff
changeset | 92 | -- (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: 
30528diff
changeset | 93 | (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: 
30528diff
changeset | 94 | 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: 
30528diff
changeset | 95 | 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: 
30528diff
changeset | 96 | 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: 
30528diff
changeset | 97 | 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: 
30528diff
changeset | 98 |            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: 
30528diff
changeset | 99 | 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: 
30528diff
changeset | 100 | 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: 
30528diff
changeset | 101 | 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: 
30528diff
changeset | 102 | "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: 
30528diff
changeset | 103 | |
| 23907 | 104 | end; | 
| 105 | ||
| 106 | (* theory setup *) | |
| 107 | ||
| 108 | 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: 
30528diff
changeset | 109 | langford_setup #> | 
| 30528 | 110 |   Attrib.setup @{binding langfordsimp} (Attrib.add_del add_simp del_simp) "Langford simpset";
 | 
| 23907 | 111 | |
| 112 | end; |