| author | wenzelm | 
| Thu, 18 Oct 2012 14:15:46 +0200 | |
| changeset 49912 | c6307ee2665d | 
| parent 49687 | 4b9034f089eb | 
| child 51510 | b4f7e6734acc | 
| permissions | -rw-r--r-- | 
| 38379 
67d71449e85b
more convenient split of class modules: class and class_declaration
 haftmann parents: 
38377diff
changeset | 1 | (* Title: Pure/Isar/class.ML | 
| 24218 | 2 | Author: Florian Haftmann, TU Muenchen | 
| 3 | ||
| 38379 
67d71449e85b
more convenient split of class modules: class and class_declaration
 haftmann parents: 
38377diff
changeset | 4 | Type classes derived from primitive axclasses and locales. | 
| 24218 | 5 | *) | 
| 6 | ||
| 38379 
67d71449e85b
more convenient split of class modules: class and class_declaration
 haftmann parents: 
38377diff
changeset | 7 | signature CLASS = | 
| 24218 | 8 | sig | 
| 25462 | 9 | (*classes*) | 
| 29526 | 10 | val is_class: theory -> class -> bool | 
| 11 | val these_params: theory -> sort -> (string * (class * (string * typ))) list | |
| 38379 
67d71449e85b
more convenient split of class modules: class and class_declaration
 haftmann parents: 
38377diff
changeset | 12 | val base_sort: theory -> class -> sort | 
| 
67d71449e85b
more convenient split of class modules: class and class_declaration
 haftmann parents: 
38377diff
changeset | 13 | val rules: theory -> class -> thm option * thm | 
| 
67d71449e85b
more convenient split of class modules: class and class_declaration
 haftmann parents: 
38377diff
changeset | 14 | val these_defs: theory -> sort -> thm list | 
| 
67d71449e85b
more convenient split of class modules: class and class_declaration
 haftmann parents: 
38377diff
changeset | 15 | val these_operations: theory -> sort | 
| 
67d71449e85b
more convenient split of class modules: class and class_declaration
 haftmann parents: 
38377diff
changeset | 16 | -> (string * (class * (typ * term))) list | 
| 42359 | 17 | val print_classes: Proof.context -> unit | 
| 25311 | 18 | val init: class -> theory -> Proof.context | 
| 38379 
67d71449e85b
more convenient split of class modules: class and class_declaration
 haftmann parents: 
38377diff
changeset | 19 | val begin: class list -> sort -> Proof.context -> Proof.context | 
| 48102 
9ed089bcad93
class target handles additional non-class term parameters appropriately
 haftmann parents: 
47289diff
changeset | 20 | val const: class -> (binding * mixfix) * (term list * term list * term) -> local_theory -> local_theory | 
| 38619 
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
 haftmann parents: 
38392diff
changeset | 21 | val abbrev: class -> Syntax.mode -> (binding * mixfix) * term -> local_theory -> local_theory | 
| 38379 
67d71449e85b
more convenient split of class modules: class and class_declaration
 haftmann parents: 
38377diff
changeset | 22 | val redeclare_operations: theory -> sort -> Proof.context -> Proof.context | 
| 29526 | 23 | val class_prefix: string -> string | 
| 38379 
67d71449e85b
more convenient split of class modules: class and class_declaration
 haftmann parents: 
38377diff
changeset | 24 | val register: class -> class list -> ((string * typ) * (string * typ)) list | 
| 
67d71449e85b
more convenient split of class modules: class and class_declaration
 haftmann parents: 
38377diff
changeset | 25 | -> sort -> morphism -> morphism -> thm option -> thm option -> thm | 
| 
67d71449e85b
more convenient split of class modules: class and class_declaration
 haftmann parents: 
38377diff
changeset | 26 | -> theory -> theory | 
| 25485 | 27 | |
| 25462 | 28 | (*instances*) | 
| 38379 
67d71449e85b
more convenient split of class modules: class and class_declaration
 haftmann parents: 
38377diff
changeset | 29 | val instantiation: string list * (string * sort) list * sort -> theory -> local_theory | 
| 26247 | 30 | val instantiation_instance: (local_theory -> local_theory) | 
| 31 | -> local_theory -> Proof.state | |
| 32 | val prove_instantiation_instance: (Proof.context -> tactic) | |
| 33 | -> local_theory -> local_theory | |
| 28666 | 34 | val prove_instantiation_exit: (Proof.context -> tactic) | 
| 35 | -> local_theory -> theory | |
| 36 | val prove_instantiation_exit_result: (morphism -> 'a -> 'b) | |
| 37 | -> (Proof.context -> 'b -> tactic) -> 'a -> local_theory -> 'b * theory | |
| 31869 | 38 | val read_multi_arity: theory -> xstring list * xstring list * xstring | 
| 39 | -> string list * (string * sort) list * sort | |
| 38348 
cf7b2121ad9d
moved instantiation target formally to class_target.ML
 haftmann parents: 
38107diff
changeset | 40 | val instantiation_cmd: xstring list * xstring list * xstring -> theory -> local_theory | 
| 38377 | 41 | val instance_arity_cmd: xstring list * xstring list * xstring -> theory -> Proof.state | 
| 25485 | 42 | |
| 31635 | 43 | (*subclasses*) | 
| 44 | val classrel: class * class -> theory -> Proof.state | |
| 45 | val classrel_cmd: xstring * xstring -> theory -> Proof.state | |
| 38379 
67d71449e85b
more convenient split of class modules: class and class_declaration
 haftmann parents: 
38377diff
changeset | 46 | val register_subclass: class * class -> morphism option -> Element.witness option | 
| 
67d71449e85b
more convenient split of class modules: class and class_declaration
 haftmann parents: 
38377diff
changeset | 47 | -> morphism -> theory -> theory | 
| 31635 | 48 | |
| 49 | (*tactics*) | |
| 29526 | 50 | val intro_classes_tac: thm list -> tactic | 
| 51 | val default_intro_tac: Proof.context -> thm list -> tactic | |
| 24218 | 52 | end; | 
| 53 | ||
| 38379 
67d71449e85b
more convenient split of class modules: class and class_declaration
 haftmann parents: 
38377diff
changeset | 54 | structure Class: CLASS = | 
| 24218 | 55 | struct | 
| 56 | ||
| 24589 | 57 | (** class data **) | 
| 24218 | 58 | |
| 46919 | 59 | datatype class_data = Class_Data of {
 | 
| 28715 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 60 | |
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 61 | (* static part *) | 
| 24218 | 62 | consts: (string * string) list | 
| 24836 | 63 | (*locale parameter ~> constant name*), | 
| 25062 | 64 | base_sort: sort, | 
| 29545 | 65 | base_morph: morphism | 
| 29439 | 66 | (*static part of canonical morphism*), | 
| 32850 | 67 | export_morph: morphism, | 
| 25618 | 68 | assm_intro: thm option, | 
| 69 | of_class: thm, | |
| 70 | axiom: thm option, | |
| 42359 | 71 | |
| 28715 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 72 | (* dynamic part *) | 
| 24657 | 73 | defs: thm list, | 
| 25368 
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
 haftmann parents: 
25344diff
changeset | 74 | operations: (string * (class * (typ * term))) list | 
| 28715 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 75 | |
| 48106 | 76 | (* n.b. | 
| 77 | params = logical parameters of class | |
| 78 | operations = operations participating in user-space type system | |
| 79 | *) | |
| 24657 | 80 | }; | 
| 24218 | 81 | |
| 32850 | 82 | fun make_class_data ((consts, base_sort, base_morph, export_morph, assm_intro, of_class, axiom), | 
| 25368 
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
 haftmann parents: 
25344diff
changeset | 83 | (defs, operations)) = | 
| 46919 | 84 |   Class_Data {consts = consts, base_sort = base_sort,
 | 
| 32850 | 85 | base_morph = base_morph, export_morph = export_morph, assm_intro = assm_intro, | 
| 46919 | 86 | of_class = of_class, axiom = axiom, defs = defs, operations = operations}; | 
| 87 | fun map_class_data f (Class_Data {consts, base_sort, base_morph, export_morph, assm_intro,
 | |
| 88 | of_class, axiom, defs, operations}) = | |
| 32850 | 89 | make_class_data (f ((consts, base_sort, base_morph, export_morph, assm_intro, of_class, axiom), | 
| 25368 
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
 haftmann parents: 
25344diff
changeset | 90 | (defs, operations))); | 
| 46919 | 91 | fun merge_class_data _ (Class_Data {consts = consts,
 | 
| 32850 | 92 | base_sort = base_sort, base_morph = base_morph, export_morph = export_morph, assm_intro = assm_intro, | 
| 46919 | 93 | of_class = of_class, axiom = axiom, defs = defs1, operations = operations1}, | 
| 94 |   Class_Data {consts = _, base_sort = _, base_morph = _, export_morph = _, assm_intro = _,
 | |
| 95 | of_class = _, axiom = _, defs = defs2, operations = operations2}) = | |
| 32850 | 96 | make_class_data ((consts, base_sort, base_morph, export_morph, assm_intro, of_class, axiom), | 
| 24914 
95cda5dd58d5
added proper subclass concept; improved class target
 haftmann parents: 
24901diff
changeset | 97 | (Thm.merge_thms (defs1, defs2), | 
| 25368 
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
 haftmann parents: 
25344diff
changeset | 98 | AList.merge (op =) (K true) (operations1, operations2))); | 
| 24218 | 99 | |
| 46919 | 100 | structure Class_Data = Theory_Data | 
| 24218 | 101 | ( | 
| 25038 | 102 | type T = class_data Graph.T | 
| 103 | val empty = Graph.empty; | |
| 24218 | 104 | val extend = I; | 
| 33522 | 105 | val merge = Graph.join merge_class_data; | 
| 24218 | 106 | ); | 
| 107 | ||
| 108 | ||
| 109 | (* queries *) | |
| 110 | ||
| 46919 | 111 | fun lookup_class_data thy class = | 
| 112 | (case try (Graph.get_node (Class_Data.get thy)) class of | |
| 113 | SOME (Class_Data data) => SOME data | |
| 114 | | NONE => NONE); | |
| 24218 | 115 | |
| 46919 | 116 | fun the_class_data thy class = | 
| 117 | (case lookup_class_data thy class of | |
| 118 |     NONE => error ("Undeclared class " ^ quote class)
 | |
| 119 | | SOME data => data); | |
| 24218 | 120 | |
| 25038 | 121 | val is_class = is_some oo lookup_class_data; | 
| 122 | ||
| 46919 | 123 | val ancestry = Graph.all_succs o Class_Data.get; | 
| 124 | val heritage = Graph.all_preds o Class_Data.get; | |
| 29509 
1ff0f3f08a7b
migrated class package to new locale implementation
 haftmann parents: 
29439diff
changeset | 125 | |
| 25002 | 126 | fun these_params thy = | 
| 24218 | 127 | let | 
| 128 | fun params class = | |
| 129 | let | |
| 24930 
cc2e0e8c81af
renamed AxClass.get_definition to AxClass.get_info (again);
 wenzelm parents: 
24920diff
changeset | 130 | val const_typs = (#params o AxClass.get_info thy) class; | 
| 24657 | 131 | val const_names = (#consts o the_class_data thy) class; | 
| 24218 | 132 | in | 
| 26518 | 133 | (map o apsnd) | 
| 134 | (fn c => (class, (c, (the o AList.lookup (op =) const_typs) c))) const_names | |
| 24218 | 135 | end; | 
| 136 | in maps params o ancestry thy end; | |
| 137 | ||
| 29526 | 138 | val base_sort = #base_sort oo the_class_data; | 
| 139 | ||
| 140 | fun rules thy class = | |
| 46919 | 141 |   let val {axiom, of_class, ...} = the_class_data thy class
 | 
| 29526 | 142 | in (axiom, of_class) end; | 
| 143 | ||
| 144 | fun all_assm_intros thy = | |
| 46919 | 145 |   Graph.fold (fn (_, (Class_Data {assm_intro, ...}, _)) => fold (insert Thm.eq_thm)
 | 
| 146 | (the_list assm_intro)) (Class_Data.get thy) []; | |
| 24218 | 147 | |
| 29526 | 148 | fun these_defs thy = maps (#defs o the_class_data thy) o ancestry thy; | 
| 149 | fun these_operations thy = maps (#operations o the_class_data thy) o ancestry thy; | |
| 29358 | 150 | |
| 29526 | 151 | val base_morphism = #base_morph oo the_class_data; | 
| 47078 | 152 | |
| 46919 | 153 | fun morphism thy class = | 
| 154 | (case Element.eq_morphism thy (these_defs thy [class]) of | |
| 155 | SOME eq_morph => base_morphism thy class $> eq_morph | |
| 156 | | NONE => base_morphism thy class); | |
| 47078 | 157 | |
| 32850 | 158 | val export_morphism = #export_morph oo the_class_data; | 
| 28715 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 159 | |
| 42359 | 160 | fun print_classes ctxt = | 
| 24218 | 161 | let | 
| 42360 | 162 | val thy = Proof_Context.theory_of ctxt; | 
| 24218 | 163 | val algebra = Sign.classes_of thy; | 
| 164 | val arities = | |
| 165 | Symtab.empty | |
| 166 | |> Symtab.fold (fn (tyco, arities) => fold (fn (class, _) => | |
| 167 | Symtab.map_default (class, []) (insert (op =) tyco)) arities) | |
| 36328 
4d9deabf6474
replaced Sorts.rep_algebra by slightly more abstract selectors classes_of/arities_of;
 wenzelm parents: 
36323diff
changeset | 168 | (Sorts.arities_of algebra); | 
| 24218 | 169 | val the_arities = these o Symtab.lookup arities; | 
| 170 | fun mk_arity class tyco = | |
| 171 | let | |
| 172 | val Ss = Sorts.mg_domain algebra tyco [class]; | |
| 24920 | 173 | in Syntax.pretty_arity ctxt (tyco, Ss, [class]) end; | 
| 42359 | 174 | fun mk_param (c, ty) = | 
| 42360 | 175 | Pretty.str (Proof_Context.extern_const ctxt c ^ " :: " ^ | 
| 49687 
4b9034f089eb
use explicit Type.strip_sorts_dummy to suppress sort constraints in output;
 wenzelm parents: 
48106diff
changeset | 176 | Syntax.string_of_typ ctxt (Type.strip_sorts_dummy ty)); | 
| 24218 | 177 | fun mk_entry class = (Pretty.block o Pretty.fbreaks o map_filter I) [ | 
| 42360 | 178 |       (SOME o Pretty.str) ("class " ^ Proof_Context.extern_class ctxt class ^ ":"),
 | 
| 24218 | 179 | (SOME o Pretty.block) [Pretty.str "supersort: ", | 
| 24920 | 180 | (Syntax.pretty_sort ctxt o Sign.minimize_sort thy o Sign.super_classes thy) class], | 
| 25062 | 181 | ((fn [] => NONE | ps => (SOME o Pretty.block o Pretty.fbreaks) | 
| 182 | (Pretty.str "parameters:" :: ps)) o map mk_param | |
| 24930 
cc2e0e8c81af
renamed AxClass.get_definition to AxClass.get_info (again);
 wenzelm parents: 
24920diff
changeset | 183 | o these o Option.map #params o try (AxClass.get_info thy)) class, | 
| 24218 | 184 | (SOME o Pretty.block o Pretty.breaks) [ | 
| 185 | Pretty.str "instances:", | |
| 186 | Pretty.list "" "" (map (mk_arity class) (the_arities class)) | |
| 187 | ] | |
| 188 | ] | |
| 189 | in | |
| 24589 | 190 | (Pretty.writeln o Pretty.chunks o separate (Pretty.str "") | 
| 191 | o map mk_entry o Sorts.all_classes) algebra | |
| 24218 | 192 | end; | 
| 193 | ||
| 194 | ||
| 195 | (* updaters *) | |
| 196 | ||
| 32850 | 197 | fun register class sups params base_sort base_morph export_morph | 
| 29358 | 198 | axiom assm_intro of_class thy = | 
| 25002 | 199 | let | 
| 25368 
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
 haftmann parents: 
25344diff
changeset | 200 | val operations = map (fn (v_ty as (_, ty), (c, _)) => | 
| 25683 | 201 | (c, (class, (ty, Free v_ty)))) params; | 
| 25038 | 202 | val add_class = Graph.new_node (class, | 
| 31599 | 203 | make_class_data (((map o pairself) fst params, base_sort, | 
| 32850 | 204 | base_morph, export_morph, assm_intro, of_class, axiom), ([], operations))) | 
| 29526 | 205 | #> fold (curry Graph.add_edge class) sups; | 
| 46919 | 206 | in Class_Data.map add_class thy end; | 
| 24218 | 207 | |
| 46919 | 208 | fun activate_defs class thms thy = | 
| 209 | (case Element.eq_morphism thy thms of | |
| 210 | SOME eq_morph => fold (fn cls => fn thy => | |
| 38107 | 211 | Context.theory_map (Locale.amend_registration (cls, base_morphism thy cls) | 
| 212 | (eq_morph, true) (export_morphism thy cls)) thy) (heritage thy [class]) thy | |
| 46919 | 213 | | NONE => thy); | 
| 29526 | 214 | |
| 25368 
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
 haftmann parents: 
25344diff
changeset | 215 | fun register_operation class (c, (t, some_def)) thy = | 
| 25062 | 216 | let | 
| 29358 | 217 | val base_sort = base_sort thy class; | 
| 29439 | 218 | val prep_typ = map_type_tfree | 
| 219 | (fn (v, sort) => if Name.aT = v | |
| 220 | then TFree (v, base_sort) else TVar ((v, 0), sort)); | |
| 25368 
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
 haftmann parents: 
25344diff
changeset | 221 | val t' = map_types prep_typ t; | 
| 
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
 haftmann parents: 
25344diff
changeset | 222 | val ty' = Term.fastype_of t'; | 
| 25062 | 223 | in | 
| 224 | thy | |
| 46919 | 225 | |> (Class_Data.map o Graph.map_node class o map_class_data o apsnd) | 
| 25368 
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
 haftmann parents: 
25344diff
changeset | 226 | (fn (defs, operations) => | 
| 25096 | 227 | (fold cons (the_list some_def) defs, | 
| 25368 
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
 haftmann parents: 
25344diff
changeset | 228 | (c, (class, (ty', t'))) :: operations)) | 
| 36674 
d95f39448121
eq_morphism is always optional: avoid trivial morphism for empty list of equations
 haftmann parents: 
36672diff
changeset | 229 | |> activate_defs class (the_list some_def) | 
| 25062 | 230 | end; | 
| 24218 | 231 | |
| 29558 | 232 | fun register_subclass (sub, sup) some_dep_morph some_wit export thy = | 
| 25618 | 233 | let | 
| 29558 | 234 | val intros = (snd o rules thy) sup :: map_filter I | 
| 35021 
c839a4c670c6
renamed old-style Drule.standard to Drule.export_without_context, to emphasize that this is in no way a standard operation;
 wenzelm parents: 
33969diff
changeset | 235 | [Option.map (Drule.export_without_context_open o Element.conclude_witness) some_wit, | 
| 29558 | 236 | (fst o rules thy) sub]; | 
| 47078 | 237 | val classrel = | 
| 238 | Skip_Proof.prove_global thy [] [] (Logic.mk_classrel (sub, sup)) | |
| 239 | (K (EVERY (map (TRYALL o Tactic.rtac) intros))); | |
| 29509 
1ff0f3f08a7b
migrated class package to new locale implementation
 haftmann parents: 
29439diff
changeset | 240 | val diff_sort = Sign.complete_sort thy [sup] | 
| 31012 | 241 | |> subtract (op =) (Sign.complete_sort thy [sub]) | 
| 242 | |> filter (is_class thy); | |
| 46919 | 243 | val add_dependency = | 
| 244 | (case some_dep_morph of | |
| 245 | SOME dep_morph => Locale.add_dependency sub | |
| 41270 | 246 | (sup, dep_morph $> Element.satisfy_morphism (the_list some_wit)) NONE export | 
| 46919 | 247 | | NONE => I); | 
| 25618 | 248 | in | 
| 249 | thy | |
| 37246 | 250 | |> AxClass.add_classrel classrel | 
| 46919 | 251 | |> Class_Data.map (Graph.add_edge (sub, sup)) | 
| 29526 | 252 | |> activate_defs sub (these_defs thy diff_sort) | 
| 32074 | 253 | |> add_dependency | 
| 24218 | 254 | end; | 
| 255 | ||
| 256 | ||
| 24589 | 257 | (** classes and class target **) | 
| 24218 | 258 | |
| 25002 | 259 | (* class context syntax *) | 
| 24748 | 260 | |
| 35858 
0d394a82337e
handle hidden polymorphism in class target (without class target syntax, though)
 haftmann parents: 
35845diff
changeset | 261 | fun these_unchecks thy = map (fn (c, (_, (ty, t))) => (t, Const (c, ty))) | 
| 
0d394a82337e
handle hidden polymorphism in class target (without class target syntax, though)
 haftmann parents: 
35845diff
changeset | 262 | o these_operations thy; | 
| 29577 | 263 | |
| 29632 
c3d576157244
fixed reading of class specs: declare class operations in context
 haftmann parents: 
29610diff
changeset | 264 | fun redeclare_const thy c = | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30344diff
changeset | 265 | let val b = Long_Name.base_name c | 
| 29632 
c3d576157244
fixed reading of class specs: declare class operations in context
 haftmann parents: 
29610diff
changeset | 266 | in Sign.intern_const thy b = c ? Variable.declare_const (b, c) end; | 
| 
c3d576157244
fixed reading of class specs: declare class operations in context
 haftmann parents: 
29610diff
changeset | 267 | |
| 29577 | 268 | fun synchronize_class_syntax sort base_sort ctxt = | 
| 24914 
95cda5dd58d5
added proper subclass concept; improved class target
 haftmann parents: 
24901diff
changeset | 269 | let | 
| 42360 | 270 | val thy = Proof_Context.theory_of ctxt; | 
| 26596 | 271 | val algebra = Sign.classes_of thy; | 
| 29577 | 272 | val operations = these_operations thy sort; | 
| 26518 | 273 | fun subst_class_typ sort = map_type_tfree (K (TVar ((Name.aT, 0), sort))); | 
| 274 | val primary_constraints = | |
| 25368 
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
 haftmann parents: 
25344diff
changeset | 275 | (map o apsnd) (subst_class_typ base_sort o fst o snd) operations; | 
| 26518 | 276 | val secondary_constraints = | 
| 25368 
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
 haftmann parents: 
25344diff
changeset | 277 | (map o apsnd) (fn (class, (ty, _)) => subst_class_typ [class] ty) operations; | 
| 46919 | 278 | fun improve (c, ty) = | 
| 279 | (case AList.lookup (op =) primary_constraints c of | |
| 280 | SOME ty' => | |
| 281 | (case try (Type.raw_match (ty', ty)) Vartab.empty of | |
| 282 | SOME tyenv => | |
| 283 | (case Vartab.lookup tyenv (Name.aT, 0) of | |
| 284 | SOME (_, ty' as TVar (vi, sort)) => | |
| 285 | if Type_Infer.is_param vi andalso Sorts.sort_le algebra (base_sort, sort) | |
| 286 | then SOME (ty', TFree (Name.aT, base_sort)) | |
| 287 | else NONE | |
| 26238 | 288 | | _ => NONE) | 
| 289 | | NONE => NONE) | |
| 46919 | 290 | | NONE => NONE); | 
| 47005 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 wenzelm parents: 
46923diff
changeset | 291 | fun subst (c, _) = Option.map snd (AList.lookup (op =) operations c); | 
| 29577 | 292 | val unchecks = these_unchecks thy sort; | 
| 25083 | 293 | in | 
| 294 | ctxt | |
| 29632 
c3d576157244
fixed reading of class specs: declare class operations in context
 haftmann parents: 
29610diff
changeset | 295 | |> fold (redeclare_const thy o fst) primary_constraints | 
| 26518 | 296 | |> Overloading.map_improvable_syntax (K (((primary_constraints, secondary_constraints), | 
| 26730 | 297 | (((improve, subst), true), unchecks)), false)) | 
| 26518 | 298 | |> Overloading.set_primary_constraints | 
| 25083 | 299 | end; | 
| 300 | ||
| 29632 
c3d576157244
fixed reading of class specs: declare class operations in context
 haftmann parents: 
29610diff
changeset | 301 | fun redeclare_operations thy sort = | 
| 
c3d576157244
fixed reading of class specs: declare class operations in context
 haftmann parents: 
29610diff
changeset | 302 | fold (redeclare_const thy o fst) (these_operations thy sort); | 
| 
c3d576157244
fixed reading of class specs: declare class operations in context
 haftmann parents: 
29610diff
changeset | 303 | |
| 29577 | 304 | fun begin sort base_sort ctxt = | 
| 25083 | 305 | ctxt | 
| 306 | |> Variable.declare_term | |
| 307 | (Logic.mk_type (TFree (Name.aT, base_sort))) | |
| 29577 | 308 | |> synchronize_class_syntax sort base_sort | 
| 39378 
df86b1b4ce10
more precise name for activation of improveable syntax
 haftmann parents: 
39134diff
changeset | 309 | |> Overloading.activate_improvable_syntax; | 
| 24901 
d3cbf79769b9
added first version of user-space type system for class target
 haftmann parents: 
24847diff
changeset | 310 | |
| 25311 | 311 | fun init class thy = | 
| 312 | thy | |
| 29509 
1ff0f3f08a7b
migrated class package to new locale implementation
 haftmann parents: 
29439diff
changeset | 313 | |> Locale.init class | 
| 29358 | 314 | |> begin [class] (base_sort thy class); | 
| 24914 
95cda5dd58d5
added proper subclass concept; improved class target
 haftmann parents: 
24901diff
changeset | 315 | |
| 24748 | 316 | |
| 27690 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 317 | (* class target *) | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 318 | |
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30344diff
changeset | 319 | val class_prefix = Logic.const_of_class o Long_Name.base_name; | 
| 29526 | 320 | |
| 47078 | 321 | local | 
| 322 | ||
| 38619 
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
 haftmann parents: 
38392diff
changeset | 323 | fun target_extension f class lthy = | 
| 
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
 haftmann parents: 
38392diff
changeset | 324 | lthy | 
| 
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
 haftmann parents: 
38392diff
changeset | 325 | |> Local_Theory.raw_theory f | 
| 47079 
6231adc3895d
synchronize syntax uniformly for target stack and aux. context;
 wenzelm parents: 
47078diff
changeset | 326 | |> Local_Theory.map_contexts | 
| 47245 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 327 | (K (synchronize_class_syntax [class] (base_sort (Proof_Context.theory_of lthy) class))); | 
| 38619 
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
 haftmann parents: 
38392diff
changeset | 328 | |
| 48102 
9ed089bcad93
class target handles additional non-class term parameters appropriately
 haftmann parents: 
47289diff
changeset | 329 | fun target_const class ((c, mx), (type_params, term_params, dict)) thy = | 
| 27690 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 330 | let | 
| 29577 | 331 | val morph = morphism thy class; | 
| 48102 
9ed089bcad93
class target handles additional non-class term parameters appropriately
 haftmann parents: 
47289diff
changeset | 332 | val class_params = map fst (these_params thy [class]); | 
| 
9ed089bcad93
class target handles additional non-class term parameters appropriately
 haftmann parents: 
47289diff
changeset | 333 | val additional_params = | 
| 
9ed089bcad93
class target handles additional non-class term parameters appropriately
 haftmann parents: 
47289diff
changeset | 334 | subtract (fn (v, Free (w, _)) => v = w | _ => false) class_params term_params; | 
| 
9ed089bcad93
class target handles additional non-class term parameters appropriately
 haftmann parents: 
47289diff
changeset | 335 | val context_params = map (Morphism.term morph) (type_params @ additional_params); | 
| 30344 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 wenzelm parents: 
30280diff
changeset | 336 | val b = Morphism.binding morph c; | 
| 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 wenzelm parents: 
30280diff
changeset | 337 | val b_def = Morphism.binding morph (Binding.suffix_name "_dict" b); | 
| 29577 | 338 | val c' = Sign.full_name thy b; | 
| 29439 | 339 | val dict' = Morphism.term morph dict; | 
| 48102 
9ed089bcad93
class target handles additional non-class term parameters appropriately
 haftmann parents: 
47289diff
changeset | 340 | val ty' = map Term.fastype_of context_params ---> Term.fastype_of dict'; | 
| 
9ed089bcad93
class target handles additional non-class term parameters appropriately
 haftmann parents: 
47289diff
changeset | 341 | val def_eq = Logic.mk_equals (list_comb (Const (c', ty'), context_params), dict') | 
| 29577 | 342 | |> map_types Type.strip_sorts; | 
| 27690 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 343 | in | 
| 29577 | 344 | thy | 
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
42360diff
changeset | 345 | |> Sign.declare_const_global ((b, Type.strip_sorts ty'), mx) | 
| 29577 | 346 | |> snd | 
| 42375 
774df7c59508
report Name_Space.declare/define, relatively to context;
 wenzelm parents: 
42360diff
changeset | 347 | |> Thm.add_def_global false false (b_def, def_eq) | 
| 36106 
19deea200358
Thm.add_axiom/add_def: return internal name of foundational axiom;
 wenzelm parents: 
35858diff
changeset | 348 | |>> apsnd Thm.varifyT_global | 
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
39438diff
changeset | 349 | |-> (fn (_, def_thm) => Global_Theory.store_thm (b_def, def_thm) | 
| 29577 | 350 | #> snd | 
| 48102 
9ed089bcad93
class target handles additional non-class term parameters appropriately
 haftmann parents: 
47289diff
changeset | 351 | #> null context_params ? register_operation class (c', (dict', SOME (Thm.symmetric def_thm)))) | 
| 27690 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 352 | |> Sign.add_const_constraint (c', SOME ty') | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 353 | end; | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 354 | |
| 38619 
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
 haftmann parents: 
38392diff
changeset | 355 | fun target_abbrev class prmode ((c, mx), rhs) thy = | 
| 27690 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 356 | let | 
| 29577 | 357 | val morph = morphism thy class; | 
| 358 | val unchecks = these_unchecks thy [class]; | |
| 30344 
10a67c5ddddb
more uniform handling of binding in targets and derived elements;
 wenzelm parents: 
30280diff
changeset | 359 | val b = Morphism.binding morph c; | 
| 29577 | 360 | val c' = Sign.full_name thy b; | 
| 27690 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 361 | val rhs' = Pattern.rewrite_term thy unchecks [] rhs; | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 362 | val ty' = Term.fastype_of rhs'; | 
| 45344 
e209da839ff4
added Logic.varify_types_global/unvarify_types_global, which avoids somewhat expensive Term.map_types;
 wenzelm parents: 
45310diff
changeset | 363 | val rhs'' = Logic.varify_types_global rhs'; | 
| 27690 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 364 | in | 
| 29577 | 365 | thy | 
| 33173 
b8ca12f6681a
eliminated obsolete tags for types/consts -- now handled via name space, in strongly typed fashion;
 wenzelm parents: 
32970diff
changeset | 366 | |> Sign.add_abbrev (#1 prmode) (b, rhs'') | 
| 29577 | 367 | |> snd | 
| 27690 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 368 | |> Sign.add_const_constraint (c', SOME ty') | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 369 | |> Sign.notation true prmode [(Const (c', ty'), mx)] | 
| 37146 
f652333bbf8e
renamed structure PrintMode to Print_Mode, keeping the old name as legacy alias for some time;
 wenzelm parents: 
37145diff
changeset | 370 | |> not (#1 prmode = Print_Mode.input) ? register_operation class (c', (rhs', NONE)) | 
| 27690 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 371 | end; | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 372 | |
| 38619 
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
 haftmann parents: 
38392diff
changeset | 373 | in | 
| 
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
 haftmann parents: 
38392diff
changeset | 374 | |
| 
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
 haftmann parents: 
38392diff
changeset | 375 | fun const class arg = target_extension (target_const class arg) class; | 
| 
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
 haftmann parents: 
38392diff
changeset | 376 | fun abbrev class prmode arg = target_extension (target_abbrev class prmode arg) class; | 
| 
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
 haftmann parents: 
38392diff
changeset | 377 | |
| 
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
 haftmann parents: 
38392diff
changeset | 378 | end; | 
| 
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
 haftmann parents: 
38392diff
changeset | 379 | |
| 27690 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 380 | |
| 31635 | 381 | (* simple subclasses *) | 
| 382 | ||
| 383 | local | |
| 384 | ||
| 385 | fun gen_classrel mk_prop classrel thy = | |
| 386 | let | |
| 387 | fun after_qed results = | |
| 42360 | 388 | Proof_Context.background_theory ((fold o fold) AxClass.add_classrel results); | 
| 31635 | 389 | in | 
| 390 | thy | |
| 42360 | 391 | |> Proof_Context.init_global | 
| 36323 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36106diff
changeset | 392 | |> Proof.theorem NONE after_qed [[(mk_prop thy classrel, [])]] | 
| 31635 | 393 | end; | 
| 394 | ||
| 395 | in | |
| 396 | ||
| 397 | val classrel = | |
| 398 | gen_classrel (Logic.mk_classrel oo AxClass.cert_classrel); | |
| 399 | val classrel_cmd = | |
| 400 | gen_classrel (Logic.mk_classrel oo AxClass.read_classrel); | |
| 401 | ||
| 402 | end; (*local*) | |
| 403 | ||
| 404 | ||
| 25462 | 405 | (** instantiation target **) | 
| 406 | ||
| 407 | (* bookkeeping *) | |
| 408 | ||
| 409 | datatype instantiation = Instantiation of {
 | |
| 25864 | 410 | arities: string list * (string * sort) list * sort, | 
| 25462 | 411 | params: ((string * string) * (string * typ)) list | 
| 25603 | 412 | (*(instantiation parameter, type constructor), (local instantiation parameter, typ)*) | 
| 25462 | 413 | } | 
| 414 | ||
| 33519 | 415 | structure Instantiation = Proof_Data | 
| 25462 | 416 | ( | 
| 46919 | 417 | type T = instantiation; | 
| 418 |   fun init _ = Instantiation {arities = ([], [], []), params = []};
 | |
| 25462 | 419 | ); | 
| 420 | ||
| 25485 | 421 | fun mk_instantiation (arities, params) = | 
| 46919 | 422 |   Instantiation {arities = arities, params = params};
 | 
| 423 | ||
| 424 | val get_instantiation = | |
| 425 | (fn Instantiation data => data) o Instantiation.get o Local_Theory.target_of; | |
| 25462 | 426 | |
| 46919 | 427 | fun map_instantiation f = | 
| 428 | (Local_Theory.target o Instantiation.map) | |
| 429 |     (fn Instantiation {arities, params} => mk_instantiation (f (arities, params)));
 | |
| 430 | ||
| 431 | fun the_instantiation lthy = | |
| 432 | (case get_instantiation lthy of | |
| 433 |     {arities = ([], [], []), ...} => error "No instantiation target"
 | |
| 434 | | data => data); | |
| 25462 | 435 | |
| 25485 | 436 | val instantiation_params = #params o get_instantiation; | 
| 25462 | 437 | |
| 30519 
c05c0199826f
coherent binding policy with primitive target operations
 haftmann parents: 
30364diff
changeset | 438 | fun instantiation_param lthy b = instantiation_params lthy | 
| 
c05c0199826f
coherent binding policy with primitive target operations
 haftmann parents: 
30364diff
changeset | 439 | |> find_first (fn (_, (v, _)) => Binding.name_of b = v) | 
| 25462 | 440 | |> Option.map (fst o fst); | 
| 441 | ||
| 31869 | 442 | fun read_multi_arity thy (raw_tycos, raw_sorts, raw_sort) = | 
| 443 | let | |
| 42360 | 444 | val ctxt = Proof_Context.init_global thy; | 
| 445 | val all_arities = map (fn raw_tyco => Proof_Context.read_arity ctxt | |
| 31869 | 446 | (raw_tyco, raw_sorts, raw_sort)) raw_tycos; | 
| 447 | val tycos = map #1 all_arities; | |
| 448 | val (_, sorts, sort) = hd all_arities; | |
| 43329 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 wenzelm parents: 
42402diff
changeset | 449 | val vs = Name.invent_names Name.context Name.aT sorts; | 
| 31869 | 450 | in (tycos, vs, sort) end; | 
| 451 | ||
| 25462 | 452 | |
| 453 | (* syntax *) | |
| 454 | ||
| 26238 | 455 | fun synchronize_inst_syntax ctxt = | 
| 25462 | 456 | let | 
| 46919 | 457 |     val Instantiation {params, ...} = Instantiation.get ctxt;
 | 
| 31249 | 458 | |
| 33969 | 459 | val lookup_inst_param = AxClass.lookup_inst_param | 
| 42360 | 460 | (Sign.consts_of (Proof_Context.theory_of ctxt)) params; | 
| 46919 | 461 | fun subst (c, ty) = | 
| 462 | (case lookup_inst_param (c, ty) of | |
| 463 | SOME (v_ty as (_, ty)) => SOME (ty, Free v_ty) | |
| 464 | | NONE => NONE); | |
| 26238 | 465 | val unchecks = | 
| 466 | map (fn ((c, _), v_ty as (_, ty)) => (Free v_ty, Const (c, ty))) params; | |
| 467 | in | |
| 468 | ctxt | |
| 26259 | 469 | |> Overloading.map_improvable_syntax | 
| 46919 | 470 | (fn (((primary_constraints, _), (((improve, _), _), _)), _) => | 
| 471 | (((primary_constraints, []), (((improve, subst), false), unchecks)), false)) | |
| 26238 | 472 | end; | 
| 25462 | 473 | |
| 42385 
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
 wenzelm parents: 
42383diff
changeset | 474 | fun resort_terms ctxt algebra consts constraints ts = | 
| 38382 | 475 | let | 
| 42385 
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
 wenzelm parents: 
42383diff
changeset | 476 | fun matchings (Const (c_ty as (c, _))) = | 
| 
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
 wenzelm parents: 
42383diff
changeset | 477 | (case constraints c of | 
| 
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
 wenzelm parents: 
42383diff
changeset | 478 | NONE => I | 
| 
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
 wenzelm parents: 
42383diff
changeset | 479 | | SOME sorts => | 
| 
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
 wenzelm parents: 
42383diff
changeset | 480 | fold2 (curry (Sorts.meet_sort algebra)) (Consts.typargs consts c_ty) sorts) | 
| 
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
 wenzelm parents: 
42383diff
changeset | 481 | | matchings _ = I; | 
| 38382 | 482 | val tvartab = (fold o fold_aterms) matchings ts Vartab.empty | 
| 47005 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 wenzelm parents: 
46923diff
changeset | 483 | handle Sorts.CLASS_ERROR e => error (Sorts.class_error (Context.pretty ctxt) e); | 
| 38382 | 484 | val inst = map_type_tvar | 
| 485 | (fn (vi, sort) => TVar (vi, the_default sort (Vartab.lookup tvartab vi))); | |
| 42402 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42388diff
changeset | 486 | in if Vartab.is_empty tvartab then ts else (map o map_types) inst ts end; | 
| 38382 | 487 | |
| 25462 | 488 | |
| 489 | (* target *) | |
| 490 | ||
| 47078 | 491 | fun define_overloaded (c, U) v (b_def, rhs) = | 
| 492 | Local_Theory.background_theory_result (AxClass.declare_overloaded (c, U) | |
| 493 | ##>> AxClass.define_overloaded b_def (c, rhs)) | |
| 38382 | 494 | ##> (map_instantiation o apsnd) (filter_out (fn (_, (v', _)) => v' = v)) | 
| 47245 
ff1770df59b8
Local_Theory.map_contexts with explicit level indication: 0 = main target at bottom;
 wenzelm parents: 
47081diff
changeset | 495 | ##> Local_Theory.map_contexts (K synchronize_inst_syntax); | 
| 38382 | 496 | |
| 47289 | 497 | fun foundation (((b, U), mx), (b_def, rhs)) params lthy = | 
| 46919 | 498 | (case instantiation_param lthy b of | 
| 499 | SOME c => | |
| 500 |       if mx <> NoSyn then error ("Illegal mixfix syntax for overloaded constant " ^ quote c)
 | |
| 501 | else lthy |> define_overloaded (c, U) (Binding.name_of b) (b_def, rhs) | |
| 47289 | 502 | | NONE => lthy |> Generic_Target.theory_foundation (((b, U), mx), (b_def, rhs)) params); | 
| 38382 | 503 | |
| 504 | fun pretty lthy = | |
| 26518 | 505 | let | 
| 46919 | 506 |     val {arities = (tycos, vs, sort), params} = the_instantiation lthy;
 | 
| 38382 | 507 | fun pr_arity tyco = Syntax.pretty_arity lthy (tyco, map snd vs, sort); | 
| 508 | fun pr_param ((c, _), (v, ty)) = | |
| 42359 | 509 | Pretty.block (Pretty.breaks | 
| 42360 | 510 | [Pretty.str v, Pretty.str "==", Pretty.str (Proof_Context.extern_const lthy c), | 
| 42359 | 511 | Pretty.str "::", Syntax.pretty_typ lthy ty]); | 
| 46923 | 512 | in Pretty.command "instantiation" :: map pr_arity tycos @ map pr_param params end; | 
| 26518 | 513 | |
| 38382 | 514 | fun conclude lthy = | 
| 515 | let | |
| 42359 | 516 | val (tycos, vs, sort) = #arities (the_instantiation lthy); | 
| 42360 | 517 | val thy = Proof_Context.theory_of lthy; | 
| 42359 | 518 | val _ = tycos |> List.app (fn tyco => | 
| 46919 | 519 | if Sign.of_sort thy (Type (tyco, map TFree vs), sort) then () | 
| 42360 | 520 |       else error ("Missing instance proof for type " ^ quote (Proof_Context.extern_type lthy tyco)));
 | 
| 38382 | 521 | in lthy end; | 
| 522 | ||
| 523 | fun instantiation (tycos, vs, sort) thy = | |
| 25462 | 524 | let | 
| 47061 
355317493f34
clarified Local_Theory.init: avoid hardwired naming policy, discontinued odd/unused group argument (cf. 5ee13e0428d2);
 wenzelm parents: 
47005diff
changeset | 525 | val naming = Sign.naming_of thy; | 
| 
355317493f34
clarified Local_Theory.init: avoid hardwired naming policy, discontinued odd/unused group argument (cf. 5ee13e0428d2);
 wenzelm parents: 
47005diff
changeset | 526 | |
| 25536 | 527 | val _ = if null tycos then error "At least one arity must be given" else (); | 
| 31249 | 528 | val class_params = these_params thy (filter (can (AxClass.get_info thy)) sort); | 
| 29632 
c3d576157244
fixed reading of class specs: declare class operations in context
 haftmann parents: 
29610diff
changeset | 529 | fun get_param tyco (param, (_, (c, ty))) = | 
| 
c3d576157244
fixed reading of class specs: declare class operations in context
 haftmann parents: 
29610diff
changeset | 530 | if can (AxClass.param_of_inst thy) (c, tyco) | 
| 25603 | 531 | then NONE else SOME ((c, tyco), | 
| 46917 | 532 | (param ^ "_" ^ Long_Name.base_name tyco, map_atyps (K (Type (tyco, map TFree vs))) ty)); | 
| 31249 | 533 | val params = map_product get_param tycos class_params |> map_filter I; | 
| 26518 | 534 | val primary_constraints = map (apsnd | 
| 31249 | 535 | (map_atyps (K (TVar ((Name.aT, 0), [])))) o snd o snd) class_params; | 
| 26518 | 536 | val algebra = Sign.classes_of thy | 
| 47005 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 wenzelm parents: 
46923diff
changeset | 537 | |> fold (fn tyco => Sorts.add_arities (Context.pretty_global thy) | 
| 26518 | 538 | (tyco, map (fn class => (class, map snd vs)) sort)) tycos; | 
| 539 | val consts = Sign.consts_of thy; | |
| 540 | val improve_constraints = AList.lookup (op =) | |
| 31249 | 541 | (map (fn (_, (class, (c, _))) => (c, [[class]])) class_params); | 
| 42402 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42388diff
changeset | 542 | fun resort_check ctxt ts = resort_terms ctxt algebra consts improve_constraints ts; | 
| 31249 | 543 | val lookup_inst_param = AxClass.lookup_inst_param consts params; | 
| 42388 | 544 | fun improve (c, ty) = | 
| 545 | (case lookup_inst_param (c, ty) of | |
| 546 | SOME (_, ty') => if Sign.typ_instance thy (ty', ty) then SOME (ty, ty') else NONE | |
| 547 | | NONE => NONE); | |
| 25485 | 548 | in | 
| 549 | thy | |
| 32379 
a97e9caebd60
additional checkpoints avoid problems in error situations
 haftmann parents: 
32074diff
changeset | 550 | |> Theory.checkpoint | 
| 42360 | 551 | |> Proof_Context.init_global | 
| 31249 | 552 | |> Instantiation.put (mk_instantiation ((tycos, vs, sort), params)) | 
| 27281 | 553 | |> fold (Variable.declare_typ o TFree) vs | 
| 31249 | 554 | |> fold (Variable.declare_names o Free o snd) params | 
| 26259 | 555 | |> (Overloading.map_improvable_syntax o apfst) | 
| 29969 
9dbb046136d0
more precise improvement in instantiation user space type system
 haftmann parents: 
29632diff
changeset | 556 | (K ((primary_constraints, []), (((improve, K NONE), false), []))) | 
| 39378 
df86b1b4ce10
more precise name for activation of improveable syntax
 haftmann parents: 
39134diff
changeset | 557 | |> Overloading.activate_improvable_syntax | 
| 45429 | 558 | |> Context.proof_map (Syntax_Phases.term_check 0 "resorting" resort_check) | 
| 26238 | 559 | |> synchronize_inst_syntax | 
| 47061 
355317493f34
clarified Local_Theory.init: avoid hardwired naming policy, discontinued odd/unused group argument (cf. 5ee13e0428d2);
 wenzelm parents: 
47005diff
changeset | 560 | |> Local_Theory.init naming | 
| 38382 | 561 |        {define = Generic_Target.define foundation,
 | 
| 47250 
6523a21076a8
clarified Generic_Target.notes: always perform Attrib.partial_evaluation;
 wenzelm parents: 
47245diff
changeset | 562 | notes = Generic_Target.notes Generic_Target.theory_notes, | 
| 47286 
392c4cd97e5c
more uniform theory_abbrev with const_declaration;
 wenzelm parents: 
47279diff
changeset | 563 | abbrev = Generic_Target.abbrev Generic_Target.theory_abbrev, | 
| 47279 
4bab649dedf0
clarified standard_declaration vs. theory_declaration;
 wenzelm parents: 
47250diff
changeset | 564 | declaration = K Generic_Target.theory_declaration, | 
| 38382 | 565 | pretty = pretty, | 
| 566 | exit = Local_Theory.target_of o conclude} | |
| 25485 | 567 | end; | 
| 568 | ||
| 38382 | 569 | fun instantiation_cmd arities thy = | 
| 570 | instantiation (read_multi_arity thy arities) thy; | |
| 26238 | 571 | |
| 25485 | 572 | fun gen_instantiation_instance do_proof after_qed lthy = | 
| 573 | let | |
| 25864 | 574 | val (tycos, vs, sort) = (#arities o the_instantiation) lthy; | 
| 575 | val arities_proof = maps (fn tyco => Logic.mk_arities (tyco, map snd vs, sort)) tycos; | |
| 25462 | 576 | fun after_qed' results = | 
| 38757 
2b3e054ae6fc
renamed Local_Theory.theory(_result) to Local_Theory.background_theory(_result) to emphasize that this belongs to the infrastructure and is rarely appropriate in user-space tools;
 wenzelm parents: 
38756diff
changeset | 577 | Local_Theory.background_theory (fold (AxClass.add_arity o Thm.varifyT_global) results) | 
| 25462 | 578 | #> after_qed; | 
| 579 | in | |
| 580 | lthy | |
| 581 | |> do_proof after_qed' arities_proof | |
| 582 | end; | |
| 583 | ||
| 25485 | 584 | val instantiation_instance = gen_instantiation_instance (fn after_qed => fn ts => | 
| 36323 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36106diff
changeset | 585 | Proof.theorem NONE (after_qed o map the_single) (map (fn t => [(t, [])]) ts)); | 
| 25462 | 586 | |
| 25485 | 587 | fun prove_instantiation_instance tac = gen_instantiation_instance (fn after_qed => | 
| 25502 | 588 | fn ts => fn lthy => after_qed (map (fn t => Goal.prove lthy [] [] t | 
| 589 |     (fn {context, ...} => tac context)) ts) lthy) I;
 | |
| 25462 | 590 | |
| 28666 | 591 | fun prove_instantiation_exit tac = prove_instantiation_instance tac | 
| 33671 | 592 | #> Local_Theory.exit_global; | 
| 28666 | 593 | |
| 594 | fun prove_instantiation_exit_result f tac x lthy = | |
| 595 | let | |
| 42360 | 596 | val morph = Proof_Context.export_morphism lthy | 
| 597 | (Proof_Context.init_global (Proof_Context.theory_of lthy)); | |
| 29439 | 598 | val y = f morph x; | 
| 28666 | 599 | in | 
| 600 | lthy | |
| 601 | |> prove_instantiation_exit (fn ctxt => tac ctxt y) | |
| 602 | |> pair y | |
| 603 | end; | |
| 604 | ||
| 29526 | 605 | |
| 31635 | 606 | (* simplified instantiation interface with no class parameter *) | 
| 607 | ||
| 31869 | 608 | fun instance_arity_cmd raw_arities thy = | 
| 31635 | 609 | let | 
| 31869 | 610 | val (tycos, vs, sort) = read_multi_arity thy raw_arities; | 
| 611 | val sorts = map snd vs; | |
| 612 | val arities = maps (fn tyco => Logic.mk_arities (tyco, sorts, sort)) tycos; | |
| 47078 | 613 | fun after_qed results = | 
| 614 | Proof_Context.background_theory ((fold o fold) AxClass.add_arity results); | |
| 31635 | 615 | in | 
| 616 | thy | |
| 42360 | 617 | |> Proof_Context.init_global | 
| 36323 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36106diff
changeset | 618 | |> Proof.theorem NONE after_qed (map (fn t => [(t, [])]) arities) | 
| 31635 | 619 | end; | 
| 620 | ||
| 621 | ||
| 29526 | 622 | (** tactics and methods **) | 
| 623 | ||
| 624 | fun intro_classes_tac facts st = | |
| 625 | let | |
| 626 | val thy = Thm.theory_of_thm st; | |
| 627 | val classes = Sign.all_classes thy; | |
| 628 | val class_trivs = map (Thm.class_triv thy) classes; | |
| 629 | val class_intros = map_filter (try (#intro o AxClass.get_info thy)) classes; | |
| 630 | val assm_intros = all_assm_intros thy; | |
| 631 | in | |
| 632 | Method.intros_tac (class_trivs @ class_intros @ assm_intros) facts st | |
| 633 | end; | |
| 634 | ||
| 635 | fun default_intro_tac ctxt [] = | |
| 39438 
c5ece2a7a86e
Isar "default" step needs to fail for solved problems, for clear distinction of '.' and '..' for example -- amending lapse introduced in 9de4d64eee3b (April 2004);
 wenzelm parents: 
39378diff
changeset | 636 | COND Thm.no_prems no_tac | 
| 
c5ece2a7a86e
Isar "default" step needs to fail for solved problems, for clear distinction of '.' and '..' for example -- amending lapse introduced in 9de4d64eee3b (April 2004);
 wenzelm parents: 
39378diff
changeset | 637 | (intro_classes_tac [] ORELSE Locale.intro_locales_tac true ctxt []) | 
| 29526 | 638 | | default_intro_tac _ _ = no_tac; | 
| 639 | ||
| 640 | fun default_tac rules ctxt facts = | |
| 641 | HEADGOAL (Method.some_rule_tac rules ctxt facts) ORELSE | |
| 642 | default_intro_tac ctxt facts; | |
| 643 | ||
| 644 | val _ = Context.>> (Context.map_theory | |
| 30515 | 645 | (Method.setup (Binding.name "intro_classes") (Scan.succeed (K (METHOD intro_classes_tac))) | 
| 646 | "back-chain introduction rules of classes" #> | |
| 647 | Method.setup (Binding.name "default") (Attrib.thms >> (METHOD oo default_tac)) | |
| 648 | "apply some intro/elim rule")); | |
| 29526 | 649 | |
| 24218 | 650 | end; | 
| 25683 | 651 |