| author | wenzelm | 
| Mon, 22 Jul 2019 16:15:40 +0200 | |
| changeset 70396 | 425c5f9bc61a | 
| parent 69592 | a80d8ec6c998 | 
| child 74561 | 8e6c973003c8 | 
| permissions | -rw-r--r-- | 
| 58657 | 1 | (* Title: Pure/Tools/plugin.ML | 
| 2 | Author: Makarius | |
| 3 | Author: Jasmin Blanchette | |
| 4 | ||
| 5 | Named plugins for definitional packages. | |
| 6 | *) | |
| 7 | ||
| 8 | (** plugin name **) | |
| 9 | ||
| 10 | signature PLUGIN_NAME = | |
| 11 | sig | |
| 12 | val check: Proof.context -> xstring * Position.T -> string | |
| 58658 
9002fe021e2f
support for indirections (defined entries), e.g. relevant for quickcheck and its sub-plugins;
 wenzelm parents: 
58657diff
changeset | 13 | val define: binding -> string list -> theory -> string * theory | 
| 
9002fe021e2f
support for indirections (defined entries), e.g. relevant for quickcheck and its sub-plugins;
 wenzelm parents: 
58657diff
changeset | 14 | val define_setup: binding -> string list -> string | 
| 58657 | 15 | val declare: binding -> theory -> string * theory | 
| 58658 
9002fe021e2f
support for indirections (defined entries), e.g. relevant for quickcheck and its sub-plugins;
 wenzelm parents: 
58657diff
changeset | 16 | val declare_setup: binding -> string | 
| 58657 | 17 | type filter = string -> bool | 
| 58660 | 18 | val default_filter: filter | 
| 19 | val make_filter: Proof.context -> (Proof.context -> filter) -> filter | |
| 58657 | 20 | val parse_filter: (Proof.context -> filter) parser | 
| 21 | end; | |
| 22 | ||
| 23 | structure Plugin_Name: PLUGIN_NAME = | |
| 24 | struct | |
| 25 | ||
| 26 | (* theory data *) | |
| 27 | ||
| 28 | structure Data = Theory_Data | |
| 29 | ( | |
| 58658 
9002fe021e2f
support for indirections (defined entries), e.g. relevant for quickcheck and its sub-plugins;
 wenzelm parents: 
58657diff
changeset | 30 | type T = string list Name_Space.table; | 
| 58657 | 31 | val empty: T = Name_Space.empty_table "plugin"; | 
| 32 | val extend = I; | |
| 33 | val merge = Name_Space.merge_tables; | |
| 34 | ); | |
| 35 | ||
| 36 | ||
| 37 | (* check *) | |
| 38 | ||
| 39 | fun check ctxt = | |
| 40 | #1 o Name_Space.check (Context.Proof ctxt) (Data.get (Proof_Context.theory_of ctxt)); | |
| 41 | ||
| 42 | val _ = Theory.setup | |
| 69592 
a80d8ec6c998
support for isabelle update -u control_cartouches;
 wenzelm parents: 
69349diff
changeset | 43 | (ML_Antiquotation.inline_embedded \<^binding>\<open>plugin\<close> | 
| 69349 
7cef9e386ffe
more accurate positions for "name" (quoted string) and "embedded" (cartouche): refer to content without delimiters, which is e.g. relevant for systematic selection/renaming of scope groups;
 wenzelm parents: 
67147diff
changeset | 44 | (Args.context -- Scan.lift Args.embedded_position >> (ML_Syntax.print_string o uncurry check))); | 
| 58657 | 45 | |
| 46 | ||
| 58658 
9002fe021e2f
support for indirections (defined entries), e.g. relevant for quickcheck and its sub-plugins;
 wenzelm parents: 
58657diff
changeset | 47 | (* indirections *) | 
| 58657 | 48 | |
| 58658 
9002fe021e2f
support for indirections (defined entries), e.g. relevant for quickcheck and its sub-plugins;
 wenzelm parents: 
58657diff
changeset | 49 | fun resolve thy = maps (fn name => | 
| 
9002fe021e2f
support for indirections (defined entries), e.g. relevant for quickcheck and its sub-plugins;
 wenzelm parents: 
58657diff
changeset | 50 | (case Name_Space.get (Data.get thy) name of | 
| 
9002fe021e2f
support for indirections (defined entries), e.g. relevant for quickcheck and its sub-plugins;
 wenzelm parents: 
58657diff
changeset | 51 | [] => [name] | 
| 
9002fe021e2f
support for indirections (defined entries), e.g. relevant for quickcheck and its sub-plugins;
 wenzelm parents: 
58657diff
changeset | 52 | | names => resolve thy names)); | 
| 
9002fe021e2f
support for indirections (defined entries), e.g. relevant for quickcheck and its sub-plugins;
 wenzelm parents: 
58657diff
changeset | 53 | |
| 
9002fe021e2f
support for indirections (defined entries), e.g. relevant for quickcheck and its sub-plugins;
 wenzelm parents: 
58657diff
changeset | 54 | fun define binding rhs thy = | 
| 58657 | 55 | let | 
| 58668 | 56 | val context = Context.Theory thy |> Name_Space.map_naming (K Name_Space.global_naming); | 
| 58658 
9002fe021e2f
support for indirections (defined entries), e.g. relevant for quickcheck and its sub-plugins;
 wenzelm parents: 
58657diff
changeset | 57 | val (name, data') = Name_Space.define context true (binding, resolve thy rhs) (Data.get thy); | 
| 58657 | 58 | val thy' = Data.put data' thy; | 
| 59 | in (name, thy') end; | |
| 60 | ||
| 58658 
9002fe021e2f
support for indirections (defined entries), e.g. relevant for quickcheck and its sub-plugins;
 wenzelm parents: 
58657diff
changeset | 61 | fun define_setup binding rhs = Context.>>> (Context.map_theory_result (define binding rhs)); | 
| 
9002fe021e2f
support for indirections (defined entries), e.g. relevant for quickcheck and its sub-plugins;
 wenzelm parents: 
58657diff
changeset | 62 | |
| 
9002fe021e2f
support for indirections (defined entries), e.g. relevant for quickcheck and its sub-plugins;
 wenzelm parents: 
58657diff
changeset | 63 | |
| 
9002fe021e2f
support for indirections (defined entries), e.g. relevant for quickcheck and its sub-plugins;
 wenzelm parents: 
58657diff
changeset | 64 | (* immediate entries *) | 
| 
9002fe021e2f
support for indirections (defined entries), e.g. relevant for quickcheck and its sub-plugins;
 wenzelm parents: 
58657diff
changeset | 65 | |
| 
9002fe021e2f
support for indirections (defined entries), e.g. relevant for quickcheck and its sub-plugins;
 wenzelm parents: 
58657diff
changeset | 66 | fun declare binding thy = define binding [] thy; | 
| 
9002fe021e2f
support for indirections (defined entries), e.g. relevant for quickcheck and its sub-plugins;
 wenzelm parents: 
58657diff
changeset | 67 | fun declare_setup binding = Context.>>> (Context.map_theory_result (declare binding)); | 
| 58657 | 68 | |
| 69 | ||
| 70 | (* filter *) | |
| 71 | ||
| 72 | type filter = string -> bool; | |
| 73 | ||
| 58660 | 74 | val default_filter: filter = K true; | 
| 75 | ||
| 58657 | 76 | fun make_filter (ctxt: Proof.context) f : filter = f ctxt; | 
| 77 | ||
| 78 | val parse_filter = | |
| 79 | Parse.position (Parse.reserved "plugins") -- | |
| 80 | Parse.position (Parse.reserved "only" >> K I || Parse.reserved "del" >> K not) -- | |
| 69349 
7cef9e386ffe
more accurate positions for "name" (quoted string) and "embedded" (cartouche): refer to content without delimiters, which is e.g. relevant for systematic selection/renaming of scope groups;
 wenzelm parents: 
67147diff
changeset | 81 | (Parse.$$$ ":" |-- Scan.repeat Parse.name_position) >> | 
| 58657 | 82 | (fn (((_, pos1), (modif, pos2)), args) => fn ctxt => | 
| 83 | let | |
| 58658 
9002fe021e2f
support for indirections (defined entries), e.g. relevant for quickcheck and its sub-plugins;
 wenzelm parents: 
58657diff
changeset | 84 | val thy = Proof_Context.theory_of ctxt; | 
| 58657 | 85 | val _ = Context_Position.reports ctxt (map (rpair Markup.quasi_keyword) [pos1, pos2]); | 
| 58658 
9002fe021e2f
support for indirections (defined entries), e.g. relevant for quickcheck and its sub-plugins;
 wenzelm parents: 
58657diff
changeset | 86 | val basic_names = resolve thy (map (check ctxt) args); | 
| 
9002fe021e2f
support for indirections (defined entries), e.g. relevant for quickcheck and its sub-plugins;
 wenzelm parents: 
58657diff
changeset | 87 | in modif o member (op =) basic_names end); | 
| 58657 | 88 | |
| 89 | end; | |
| 90 | ||
| 91 | ||
| 92 | ||
| 93 | (** plugin content **) | |
| 94 | ||
| 95 | signature PLUGIN = | |
| 96 | sig | |
| 97 | type T | |
| 98 | val data: Plugin_Name.filter -> T -> local_theory -> local_theory | |
| 58666 | 99 | val data_default: T -> local_theory -> local_theory | 
| 58657 | 100 | val interpretation: string -> (T -> local_theory -> local_theory) -> theory -> theory | 
| 101 | end; | |
| 102 | ||
| 103 | functor Plugin(type T): PLUGIN = | |
| 104 | struct | |
| 105 | ||
| 106 | type T = T; | |
| 107 | ||
| 108 | ||
| 109 | (* data with interpretation *) | |
| 110 | ||
| 111 | type data = {filter: Plugin_Name.filter, naming: Name_Space.naming, value: T, id: serial};
 | |
| 112 | type interp = {name: string, apply: T -> local_theory -> local_theory, id: serial};
 | |
| 113 | ||
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
58668diff
changeset | 114 | val eq_data: data * data -> bool = op = o apply2 #id; | 
| 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
58668diff
changeset | 115 | val eq_interp: interp * interp -> bool = op = o apply2 #id; | 
| 58657 | 116 | |
| 117 | fun mk_data filter naming x : data = {filter = filter, naming = naming, value = x, id = serial ()};
 | |
| 118 | fun mk_interp name f : interp = {name = name, apply = f, id = serial ()};
 | |
| 119 | ||
| 120 | ||
| 121 | (* theory data *) | |
| 122 | ||
| 123 | structure Plugin_Data = Theory_Data | |
| 124 | ( | |
| 125 | type T = data list * (interp * data list) list; | |
| 126 | val empty : T = ([], []); | |
| 127 | val extend = I; | |
| 128 | val merge_data = merge eq_data; | |
| 129 | fun merge ((data1, interps1), (data2, interps2)) : T = | |
| 130 | (merge_data (data1, data2), AList.join eq_interp (K merge_data) (interps1, interps2)); | |
| 131 | ); | |
| 132 | ||
| 133 | ||
| 134 | (* consolidate data wrt. interpretations *) | |
| 135 | ||
| 136 | local | |
| 137 | ||
| 138 | fun apply change_naming (interp: interp) (data: data) lthy = | |
| 139 | lthy | |
| 59880 
30687c3f2b10
clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
 wenzelm parents: 
59145diff
changeset | 140 | |> change_naming ? Local_Theory.map_background_naming (K (#naming data)) | 
| 58657 | 141 | |> #apply interp (#value data) | 
| 59880 
30687c3f2b10
clarified role of naming for background theory: transform_binding (e.g. for "concealed" flag) uses naming of hypothetical context;
 wenzelm parents: 
59145diff
changeset | 142 | |> Local_Theory.restore_background_naming lthy; | 
| 58657 | 143 | |
| 144 | fun unfinished data (interp: interp, data') = | |
| 145 | (interp, | |
| 146 | if eq_list eq_data (data, data') then [] | |
| 147 | else data |> filter (fn d => #filter d (#name interp) andalso not (member eq_data data' d))); | |
| 148 | ||
| 149 | fun unfinished_data thy = | |
| 150 | let | |
| 151 | val (data, interps) = Plugin_Data.get thy; | |
| 152 | val finished = map (apsnd (K data)) interps; | |
| 153 | val thy' = Plugin_Data.put (data, finished) thy; | |
| 154 | in (map (unfinished data) interps, thy') end; | |
| 155 | ||
| 156 | in | |
| 157 | ||
| 158 | val consolidate = | |
| 159 | Local_Theory.raw_theory_result unfinished_data | |
| 160 | #-> fold_rev (fn (interp, data) => fold_rev (apply false interp) data); | |
| 161 | ||
| 162 | val consolidate' = | |
| 163 | unfinished_data #> (fn (unfinished, thy) => | |
| 164 | if forall (null o #2) unfinished then NONE | |
| 165 | else | |
| 58665 | 166 | SOME (Named_Target.theory_map | 
| 167 | (fold_rev (fn (interp, data) => fold_rev (apply true interp) data) unfinished) thy)); | |
| 58657 | 168 | |
| 169 | end; | |
| 170 | ||
| 171 | val _ = Theory.setup (Theory.at_begin consolidate'); | |
| 172 | ||
| 173 | ||
| 174 | (* add content *) | |
| 175 | ||
| 176 | fun data filter x = | |
| 177 | Local_Theory.background_theory (fn thy => | |
| 178 | Plugin_Data.map (apfst (cons (mk_data filter (Sign.naming_of thy) x))) thy) | |
| 179 | #> consolidate; | |
| 180 | ||
| 58666 | 181 | val data_default = data Plugin_Name.default_filter; | 
| 182 | ||
| 58657 | 183 | fun interpretation name f = | 
| 184 | Plugin_Data.map (apsnd (cons (mk_interp name f, []))) | |
| 185 | #> perhaps consolidate'; | |
| 186 | ||
| 187 | end; |