| author | wenzelm | 
| Thu, 02 Jul 2015 12:39:08 +0200 | |
| changeset 60630 | fc7625ec7427 | 
| parent 60619 | 7512716f03db | 
| child 60816 | 92913f915e3d | 
| 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 | 
| 60346 | 16 | -> (string * (class * ((typ * term) * bool))) 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 | 
| 57067 | 20 | val const: class -> (binding * mixfix) * term -> term list * term list -> local_theory -> local_theory | 
| 60344 
a40369ea3ba5
self-contained formulation of abbrev for named targets
 haftmann parents: 
60341diff
changeset | 21 | val abbrev: class -> Syntax.mode -> (binding * mixfix) * term -> local_theory -> (term * term) * 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 | 
| 54866 
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
 haftmann parents: 
54742diff
changeset | 47 | -> morphism -> local_theory -> local_theory | 
| 31635 | 48 | |
| 49 | (*tactics*) | |
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59423diff
changeset | 50 | val intro_classes_tac: Proof.context -> thm list -> tactic | 
| 60618 
4c79543cc376
renamed "default" to "standard", to make semantically clear what it is;
 wenzelm parents: 
60347diff
changeset | 51 | val standard_intro_classes_tac: Proof.context -> thm list -> tactic | 
| 59296 
002d817b4c37
formal pretty bodies for class specifications, accepting additional formal bookkeeping in locale.ML
 haftmann parents: 
59150diff
changeset | 52 | |
| 
002d817b4c37
formal pretty bodies for class specifications, accepting additional formal bookkeeping in locale.ML
 haftmann parents: 
59150diff
changeset | 53 | (*diagnostics*) | 
| 59423 
3a0044e95eba
backed out obsolete workaround from ef1edfb36af7
 haftmann parents: 
59420diff
changeset | 54 | val pretty_specification: theory -> class -> Pretty.T list | 
| 24218 | 55 | end; | 
| 56 | ||
| 38379 
67d71449e85b
more convenient split of class modules: class and class_declaration
 haftmann parents: 
38377diff
changeset | 57 | structure Class: CLASS = | 
| 24218 | 58 | struct | 
| 59 | ||
| 24589 | 60 | (** class data **) | 
| 24218 | 61 | |
| 46919 | 62 | datatype class_data = Class_Data of {
 | 
| 28715 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 63 | |
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 64 | (* static part *) | 
| 24218 | 65 | consts: (string * string) list | 
| 24836 | 66 | (*locale parameter ~> constant name*), | 
| 25062 | 67 | base_sort: sort, | 
| 29545 | 68 | base_morph: morphism | 
| 29439 | 69 | (*static part of canonical morphism*), | 
| 32850 | 70 | export_morph: morphism, | 
| 25618 | 71 | assm_intro: thm option, | 
| 72 | of_class: thm, | |
| 73 | axiom: thm option, | |
| 42359 | 74 | |
| 28715 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 75 | (* dynamic part *) | 
| 24657 | 76 | defs: thm list, | 
| 60346 | 77 | operations: (string * (class * ((typ * term) * bool))) list | 
| 28715 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 78 | |
| 48106 | 79 | (* n.b. | 
| 80 | params = logical parameters of class | |
| 81 | operations = operations participating in user-space type system | |
| 82 | *) | |
| 24657 | 83 | }; | 
| 24218 | 84 | |
| 32850 | 85 | 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 | 86 | (defs, operations)) = | 
| 46919 | 87 |   Class_Data {consts = consts, base_sort = base_sort,
 | 
| 32850 | 88 | base_morph = base_morph, export_morph = export_morph, assm_intro = assm_intro, | 
| 46919 | 89 | of_class = of_class, axiom = axiom, defs = defs, operations = operations}; | 
| 90 | fun map_class_data f (Class_Data {consts, base_sort, base_morph, export_morph, assm_intro,
 | |
| 91 | of_class, axiom, defs, operations}) = | |
| 32850 | 92 | 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 | 93 | (defs, operations))); | 
| 46919 | 94 | fun merge_class_data _ (Class_Data {consts = consts,
 | 
| 32850 | 95 | base_sort = base_sort, base_morph = base_morph, export_morph = export_morph, assm_intro = assm_intro, | 
| 46919 | 96 | of_class = of_class, axiom = axiom, defs = defs1, operations = operations1}, | 
| 97 |   Class_Data {consts = _, base_sort = _, base_morph = _, export_morph = _, assm_intro = _,
 | |
| 98 | of_class = _, axiom = _, defs = defs2, operations = operations2}) = | |
| 32850 | 99 | 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 | 100 | (Thm.merge_thms (defs1, defs2), | 
| 25368 
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
 haftmann parents: 
25344diff
changeset | 101 | AList.merge (op =) (K true) (operations1, operations2))); | 
| 24218 | 102 | |
| 46919 | 103 | structure Class_Data = Theory_Data | 
| 24218 | 104 | ( | 
| 25038 | 105 | type T = class_data Graph.T | 
| 106 | val empty = Graph.empty; | |
| 24218 | 107 | val extend = I; | 
| 33522 | 108 | val merge = Graph.join merge_class_data; | 
| 24218 | 109 | ); | 
| 110 | ||
| 111 | ||
| 112 | (* queries *) | |
| 113 | ||
| 46919 | 114 | fun lookup_class_data thy class = | 
| 115 | (case try (Graph.get_node (Class_Data.get thy)) class of | |
| 116 | SOME (Class_Data data) => SOME data | |
| 117 | | NONE => NONE); | |
| 24218 | 118 | |
| 46919 | 119 | fun the_class_data thy class = | 
| 120 | (case lookup_class_data thy class of | |
| 121 |     NONE => error ("Undeclared class " ^ quote class)
 | |
| 122 | | SOME data => data); | |
| 24218 | 123 | |
| 25038 | 124 | val is_class = is_some oo lookup_class_data; | 
| 125 | ||
| 46919 | 126 | val ancestry = Graph.all_succs o Class_Data.get; | 
| 127 | val heritage = Graph.all_preds o Class_Data.get; | |
| 29509 
1ff0f3f08a7b
migrated class package to new locale implementation
 haftmann parents: 
29439diff
changeset | 128 | |
| 25002 | 129 | fun these_params thy = | 
| 24218 | 130 | let | 
| 131 | fun params class = | |
| 132 | let | |
| 51685 
385ef6706252
more standard module name Axclass (according to file name);
 wenzelm parents: 
51578diff
changeset | 133 | val const_typs = (#params o Axclass.get_info thy) class; | 
| 24657 | 134 | val const_names = (#consts o the_class_data thy) class; | 
| 24218 | 135 | in | 
| 26518 | 136 | (map o apsnd) | 
| 137 | (fn c => (class, (c, (the o AList.lookup (op =) const_typs) c))) const_names | |
| 24218 | 138 | end; | 
| 139 | in maps params o ancestry thy end; | |
| 140 | ||
| 29526 | 141 | val base_sort = #base_sort oo the_class_data; | 
| 142 | ||
| 143 | fun rules thy class = | |
| 46919 | 144 |   let val {axiom, of_class, ...} = the_class_data thy class
 | 
| 29526 | 145 | in (axiom, of_class) end; | 
| 146 | ||
| 147 | fun all_assm_intros thy = | |
| 46919 | 148 |   Graph.fold (fn (_, (Class_Data {assm_intro, ...}, _)) => fold (insert Thm.eq_thm)
 | 
| 149 | (the_list assm_intro)) (Class_Data.get thy) []; | |
| 24218 | 150 | |
| 29526 | 151 | fun these_defs thy = maps (#defs o the_class_data thy) o ancestry thy; | 
| 152 | fun these_operations thy = maps (#operations o the_class_data thy) o ancestry thy; | |
| 29358 | 153 | |
| 29526 | 154 | val base_morphism = #base_morph oo the_class_data; | 
| 47078 | 155 | |
| 46919 | 156 | fun morphism thy class = | 
| 157 | (case Element.eq_morphism thy (these_defs thy [class]) of | |
| 158 | SOME eq_morph => base_morphism thy class $> eq_morph | |
| 159 | | NONE => base_morphism thy class); | |
| 47078 | 160 | |
| 32850 | 161 | val export_morphism = #export_morph oo the_class_data; | 
| 28715 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 162 | |
| 42359 | 163 | fun print_classes ctxt = | 
| 24218 | 164 | let | 
| 42360 | 165 | val thy = Proof_Context.theory_of ctxt; | 
| 24218 | 166 | val algebra = Sign.classes_of thy; | 
| 51510 
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
 wenzelm parents: 
49687diff
changeset | 167 | |
| 
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
 wenzelm parents: 
49687diff
changeset | 168 | val class_space = Proof_Context.class_space ctxt; | 
| 
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
 wenzelm parents: 
49687diff
changeset | 169 | val type_space = Proof_Context.type_space ctxt; | 
| 
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
 wenzelm parents: 
49687diff
changeset | 170 | val const_space = Proof_Context.const_space ctxt; | 
| 
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
 wenzelm parents: 
49687diff
changeset | 171 | |
| 24218 | 172 | val arities = | 
| 173 | Symtab.empty | |
| 174 | |> Symtab.fold (fn (tyco, arities) => fold (fn (class, _) => | |
| 175 | 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 | 176 | (Sorts.arities_of algebra); | 
| 51510 
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
 wenzelm parents: 
49687diff
changeset | 177 | |
| 
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
 wenzelm parents: 
49687diff
changeset | 178 | fun prt_supersort class = | 
| 
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
 wenzelm parents: 
49687diff
changeset | 179 | Syntax.pretty_sort ctxt (Sign.minimize_sort thy (Sign.super_classes thy class)); | 
| 
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
 wenzelm parents: 
49687diff
changeset | 180 | |
| 
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
 wenzelm parents: 
49687diff
changeset | 181 | fun prt_arity class tyco = | 
| 24218 | 182 | let | 
| 183 | val Ss = Sorts.mg_domain algebra tyco [class]; | |
| 24920 | 184 | in Syntax.pretty_arity ctxt (tyco, Ss, [class]) end; | 
| 51510 
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
 wenzelm parents: 
49687diff
changeset | 185 | |
| 
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
 wenzelm parents: 
49687diff
changeset | 186 | fun prt_param (c, ty) = | 
| 
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
 wenzelm parents: 
49687diff
changeset | 187 | Pretty.block | 
| 53539 | 188 | [Name_Space.pretty ctxt const_space c, Pretty.str " ::", | 
| 51510 
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
 wenzelm parents: 
49687diff
changeset | 189 | Pretty.brk 1, Syntax.pretty_typ ctxt (Type.strip_sorts_dummy ty)]; | 
| 
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
 wenzelm parents: 
49687diff
changeset | 190 | |
| 
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
 wenzelm parents: 
49687diff
changeset | 191 | fun prt_entry class = | 
| 
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
 wenzelm parents: 
49687diff
changeset | 192 | Pretty.block | 
| 55763 | 193 | ([Pretty.keyword1 "class", Pretty.brk 1, | 
| 53539 | 194 | Name_Space.pretty ctxt class_space class, Pretty.str ":", Pretty.fbrk, | 
| 195 | Pretty.block [Pretty.str "supersort: ", prt_supersort class]] @ | |
| 51685 
385ef6706252
more standard module name Axclass (according to file name);
 wenzelm parents: 
51578diff
changeset | 196 | (case try (Axclass.get_info thy) class of | 
| 51510 
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
 wenzelm parents: 
49687diff
changeset | 197 | NONE => [] | 
| 
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
 wenzelm parents: 
49687diff
changeset | 198 |           | SOME {params, ...} =>
 | 
| 
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
 wenzelm parents: 
49687diff
changeset | 199 | [Pretty.fbrk, Pretty.big_list "parameters:" (map prt_param params)]) @ | 
| 
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
 wenzelm parents: 
49687diff
changeset | 200 | (case Symtab.lookup arities class of | 
| 
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
 wenzelm parents: 
49687diff
changeset | 201 | NONE => [] | 
| 
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
 wenzelm parents: 
49687diff
changeset | 202 | | SOME ars => | 
| 
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
 wenzelm parents: 
49687diff
changeset | 203 | [Pretty.fbrk, Pretty.big_list "instances:" | 
| 
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
 wenzelm parents: 
49687diff
changeset | 204 | (map (prt_arity class) (sort (Name_Space.extern_ord ctxt type_space) ars))])); | 
| 24218 | 205 | in | 
| 51510 
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
 wenzelm parents: 
49687diff
changeset | 206 | Sorts.all_classes algebra | 
| 
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
 wenzelm parents: 
49687diff
changeset | 207 | |> sort (Name_Space.extern_ord ctxt class_space) | 
| 
b4f7e6734acc
tuned print_classes: more standard order, markup, formatting;
 wenzelm parents: 
49687diff
changeset | 208 | |> map prt_entry | 
| 56334 
6b3739fee456
some shortcuts for chunks, which sometimes avoid bulky string output;
 wenzelm parents: 
56204diff
changeset | 209 | |> Pretty.writeln_chunks2 | 
| 24218 | 210 | end; | 
| 211 | ||
| 212 | ||
| 213 | (* updaters *) | |
| 214 | ||
| 32850 | 215 | fun register class sups params base_sort base_morph export_morph | 
| 52636 | 216 | some_axiom some_assm_intro of_class thy = | 
| 25002 | 217 | let | 
| 25368 
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
 haftmann parents: 
25344diff
changeset | 218 | val operations = map (fn (v_ty as (_, ty), (c, _)) => | 
| 60346 | 219 | (c, (class, ((ty, Free v_ty), false)))) params; | 
| 25038 | 220 | val add_class = Graph.new_node (class, | 
| 59058 
a78612c67ec0
renamed "pairself" to "apply2", in accordance to @{apply 2};
 wenzelm parents: 
58837diff
changeset | 221 | make_class_data (((map o apply2) fst params, base_sort, | 
| 52636 | 222 | base_morph, export_morph, some_assm_intro, of_class, some_axiom), ([], operations))) | 
| 29526 | 223 | #> fold (curry Graph.add_edge class) sups; | 
| 46919 | 224 | in Class_Data.map add_class thy end; | 
| 24218 | 225 | |
| 46919 | 226 | fun activate_defs class thms thy = | 
| 227 | (case Element.eq_morphism thy thms of | |
| 228 | SOME eq_morph => fold (fn cls => fn thy => | |
| 38107 | 229 | Context.theory_map (Locale.amend_registration (cls, base_morphism thy cls) | 
| 230 | (eq_morph, true) (export_morphism thy cls)) thy) (heritage thy [class]) thy | |
| 46919 | 231 | | NONE => thy); | 
| 29526 | 232 | |
| 60346 | 233 | fun register_operation class (c, t) input_only thy = | 
| 25062 | 234 | let | 
| 29358 | 235 | val base_sort = base_sort thy class; | 
| 29439 | 236 | val prep_typ = map_type_tfree | 
| 237 | (fn (v, sort) => if Name.aT = v | |
| 238 | 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 | 239 | val t' = map_types prep_typ t; | 
| 
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
 haftmann parents: 
25344diff
changeset | 240 | val ty' = Term.fastype_of t'; | 
| 25062 | 241 | in | 
| 242 | thy | |
| 57173 
897cc57a6f55
always refine interpretation morphism using canonical constant's definition theorem
 haftmann parents: 
57161diff
changeset | 243 | |> (Class_Data.map o Graph.map_node class o map_class_data o apsnd o apsnd) | 
| 60346 | 244 | (cons (c, (class, ((ty', t'), input_only)))) | 
| 57173 
897cc57a6f55
always refine interpretation morphism using canonical constant's definition theorem
 haftmann parents: 
57161diff
changeset | 245 | end; | 
| 
897cc57a6f55
always refine interpretation morphism using canonical constant's definition theorem
 haftmann parents: 
57161diff
changeset | 246 | |
| 
897cc57a6f55
always refine interpretation morphism using canonical constant's definition theorem
 haftmann parents: 
57161diff
changeset | 247 | fun register_def class def_thm thy = | 
| 
897cc57a6f55
always refine interpretation morphism using canonical constant's definition theorem
 haftmann parents: 
57161diff
changeset | 248 | let | 
| 
897cc57a6f55
always refine interpretation morphism using canonical constant's definition theorem
 haftmann parents: 
57161diff
changeset | 249 | val sym_thm = Thm.symmetric def_thm | 
| 
897cc57a6f55
always refine interpretation morphism using canonical constant's definition theorem
 haftmann parents: 
57161diff
changeset | 250 | in | 
| 
897cc57a6f55
always refine interpretation morphism using canonical constant's definition theorem
 haftmann parents: 
57161diff
changeset | 251 | thy | 
| 
897cc57a6f55
always refine interpretation morphism using canonical constant's definition theorem
 haftmann parents: 
57161diff
changeset | 252 | |> (Class_Data.map o Graph.map_node class o map_class_data o apsnd o apfst) | 
| 
897cc57a6f55
always refine interpretation morphism using canonical constant's definition theorem
 haftmann parents: 
57161diff
changeset | 253 | (cons sym_thm) | 
| 
897cc57a6f55
always refine interpretation morphism using canonical constant's definition theorem
 haftmann parents: 
57161diff
changeset | 254 | |> activate_defs class [sym_thm] | 
| 25062 | 255 | end; | 
| 24218 | 256 | |
| 257 | ||
| 24589 | 258 | (** classes and class target **) | 
| 24218 | 259 | |
| 25002 | 260 | (* class context syntax *) | 
| 24748 | 261 | |
| 60347 | 262 | fun make_rewrite t c_ty = | 
| 263 | let | |
| 264 | val (vs, t') = strip_abs t; | |
| 265 | val vts = map snd vs | |
| 266 | |> Name.invent_names Name.context Name.uu | |
| 267 | |> map (fn (v, T) => Var ((v, 0), T)); | |
| 268 | in (betapplys (t, vts), betapplys (Const c_ty, vts)) end; | |
| 269 | ||
| 60346 | 270 | fun these_unchecks thy = | 
| 60347 | 271 | these_operations thy | 
| 272 | #> map_filter (fn (c, (_, ((ty, t), input_only))) => | |
| 273 | if input_only then NONE else SOME (make_rewrite t (c, ty))); | |
| 274 | ||
| 275 | fun these_unchecks_reversed thy = | |
| 276 | these_operations thy | |
| 277 | #> map (fn (c, (_, ((ty, t), _))) => (Const (c, ty), t)); | |
| 29577 | 278 | |
| 29632 
c3d576157244
fixed reading of class specs: declare class operations in context
 haftmann parents: 
29610diff
changeset | 279 | fun redeclare_const thy c = | 
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30344diff
changeset | 280 | let val b = Long_Name.base_name c | 
| 29632 
c3d576157244
fixed reading of class specs: declare class operations in context
 haftmann parents: 
29610diff
changeset | 281 | 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 | 282 | |
| 29577 | 283 | fun synchronize_class_syntax sort base_sort ctxt = | 
| 24914 
95cda5dd58d5
added proper subclass concept; improved class target
 haftmann parents: 
24901diff
changeset | 284 | let | 
| 42360 | 285 | val thy = Proof_Context.theory_of ctxt; | 
| 26596 | 286 | val algebra = Sign.classes_of thy; | 
| 29577 | 287 | val operations = these_operations thy sort; | 
| 26518 | 288 | fun subst_class_typ sort = map_type_tfree (K (TVar ((Name.aT, 0), sort))); | 
| 289 | val primary_constraints = | |
| 60346 | 290 | (map o apsnd) (subst_class_typ base_sort o fst o fst o snd) operations; | 
| 26518 | 291 | val secondary_constraints = | 
| 60346 | 292 | (map o apsnd) (fn (class, ((ty, _), _)) => subst_class_typ [class] ty) operations; | 
| 46919 | 293 | fun improve (c, ty) = | 
| 294 | (case AList.lookup (op =) primary_constraints c of | |
| 295 | SOME ty' => | |
| 296 | (case try (Type.raw_match (ty', ty)) Vartab.empty of | |
| 297 | SOME tyenv => | |
| 298 | (case Vartab.lookup tyenv (Name.aT, 0) of | |
| 299 | SOME (_, ty' as TVar (vi, sort)) => | |
| 300 | if Type_Infer.is_param vi andalso Sorts.sort_le algebra (base_sort, sort) | |
| 301 | then SOME (ty', TFree (Name.aT, base_sort)) | |
| 302 | else NONE | |
| 26238 | 303 | | _ => NONE) | 
| 304 | | NONE => NONE) | |
| 46919 | 305 | | NONE => NONE); | 
| 60346 | 306 | fun subst (c, _) = Option.map (fst o snd) (AList.lookup (op =) operations c); | 
| 29577 | 307 | val unchecks = these_unchecks thy sort; | 
| 25083 | 308 | in | 
| 309 | ctxt | |
| 29632 
c3d576157244
fixed reading of class specs: declare class operations in context
 haftmann parents: 
29610diff
changeset | 310 | |> fold (redeclare_const thy o fst) primary_constraints | 
| 60338 | 311 |     |> Overloading.map_improvable_syntax (K {primary_constraints = primary_constraints,
 | 
| 312 | secondary_constraints = secondary_constraints, improve = improve, subst = subst, | |
| 313 | no_subst_in_abbrev_mode = true, unchecks = unchecks}) | |
| 26518 | 314 | |> Overloading.set_primary_constraints | 
| 25083 | 315 | end; | 
| 316 | ||
| 54866 
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
 haftmann parents: 
54742diff
changeset | 317 | fun synchronize_class_syntax_target class lthy = | 
| 
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
 haftmann parents: 
54742diff
changeset | 318 | lthy | 
| 
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
 haftmann parents: 
54742diff
changeset | 319 | |> Local_Theory.map_contexts | 
| 
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
 haftmann parents: 
54742diff
changeset | 320 | (K (synchronize_class_syntax [class] (base_sort (Proof_Context.theory_of lthy) class))); | 
| 
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
 haftmann parents: 
54742diff
changeset | 321 | |
| 29632 
c3d576157244
fixed reading of class specs: declare class operations in context
 haftmann parents: 
29610diff
changeset | 322 | fun redeclare_operations thy sort = | 
| 
c3d576157244
fixed reading of class specs: declare class operations in context
 haftmann parents: 
29610diff
changeset | 323 | 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 | 324 | |
| 29577 | 325 | fun begin sort base_sort ctxt = | 
| 25083 | 326 | ctxt | 
| 327 | |> Variable.declare_term | |
| 328 | (Logic.mk_type (TFree (Name.aT, base_sort))) | |
| 29577 | 329 | |> synchronize_class_syntax sort base_sort | 
| 39378 
df86b1b4ce10
more precise name for activation of improveable syntax
 haftmann parents: 
39134diff
changeset | 330 | |> Overloading.activate_improvable_syntax; | 
| 24901 
d3cbf79769b9
added first version of user-space type system for class target
 haftmann parents: 
24847diff
changeset | 331 | |
| 25311 | 332 | fun init class thy = | 
| 333 | thy | |
| 29509 
1ff0f3f08a7b
migrated class package to new locale implementation
 haftmann parents: 
29439diff
changeset | 334 | |> Locale.init class | 
| 29358 | 335 | |> begin [class] (base_sort thy class); | 
| 24914 
95cda5dd58d5
added proper subclass concept; improved class target
 haftmann parents: 
24901diff
changeset | 336 | |
| 24748 | 337 | |
| 27690 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 338 | (* class target *) | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 339 | |
| 30364 
577edc39b501
moved basic algebra of long names from structure NameSpace to Long_Name;
 wenzelm parents: 
30344diff
changeset | 340 | val class_prefix = Logic.const_of_class o Long_Name.base_name; | 
| 29526 | 341 | |
| 57187 
de00494fa8b3
less ad-hoc verision of educated guess: guess identity of declaration morphism wrt. canonical morphism rather than observing particular effects of declaration morphisms only;
 haftmann parents: 
57186diff
changeset | 342 | fun guess_morphism_identity (b, rhs) phi1 phi2 = | 
| 
de00494fa8b3
less ad-hoc verision of educated guess: guess identity of declaration morphism wrt. canonical morphism rather than observing particular effects of declaration morphisms only;
 haftmann parents: 
57186diff
changeset | 343 | let | 
| 
de00494fa8b3
less ad-hoc verision of educated guess: guess identity of declaration morphism wrt. canonical morphism rather than observing particular effects of declaration morphisms only;
 haftmann parents: 
57186diff
changeset | 344 | (*FIXME proper concept to identify morphism instead of educated guess*) | 
| 58668 | 345 | val name_of_binding = Name_Space.full_name Name_Space.global_naming; | 
| 57187 
de00494fa8b3
less ad-hoc verision of educated guess: guess identity of declaration morphism wrt. canonical morphism rather than observing particular effects of declaration morphisms only;
 haftmann parents: 
57186diff
changeset | 346 | val n1 = (name_of_binding o Morphism.binding phi1) b; | 
| 
de00494fa8b3
less ad-hoc verision of educated guess: guess identity of declaration morphism wrt. canonical morphism rather than observing particular effects of declaration morphisms only;
 haftmann parents: 
57186diff
changeset | 347 | val n2 = (name_of_binding o Morphism.binding phi2) b; | 
| 
de00494fa8b3
less ad-hoc verision of educated guess: guess identity of declaration morphism wrt. canonical morphism rather than observing particular effects of declaration morphisms only;
 haftmann parents: 
57186diff
changeset | 348 | val rhs1 = Morphism.term phi1 rhs; | 
| 
de00494fa8b3
less ad-hoc verision of educated guess: guess identity of declaration morphism wrt. canonical morphism rather than observing particular effects of declaration morphisms only;
 haftmann parents: 
57186diff
changeset | 349 | val rhs2 = Morphism.term phi2 rhs; | 
| 
de00494fa8b3
less ad-hoc verision of educated guess: guess identity of declaration morphism wrt. canonical morphism rather than observing particular effects of declaration morphisms only;
 haftmann parents: 
57186diff
changeset | 350 | in n1 = n2 andalso Term.aconv_untyped (rhs1, rhs2) end; | 
| 
de00494fa8b3
less ad-hoc verision of educated guess: guess identity of declaration morphism wrt. canonical morphism rather than observing particular effects of declaration morphisms only;
 haftmann parents: 
57186diff
changeset | 351 | |
| 60345 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 352 | fun target_const class phi0 prmode (b, rhs) lthy = | 
| 57189 
5140ddfccea7
re-unified approach towards class and locale consts, with refined terminology: foo_const_declaration denotes declaration for a particular logical layer, foo_const the full stack for a particular target
 haftmann parents: 
57188diff
changeset | 353 | let | 
| 60345 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 354 | val export = Variable.export_morphism lthy (Local_Theory.target_of lthy); | 
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 355 | val guess_identity = guess_morphism_identity (b, rhs) export; | 
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 356 | val guess_canonical = guess_morphism_identity (b, rhs) (export $> phi0); | 
| 57189 
5140ddfccea7
re-unified approach towards class and locale consts, with refined terminology: foo_const_declaration denotes declaration for a particular logical layer, foo_const the full stack for a particular target
 haftmann parents: 
57188diff
changeset | 357 | in | 
| 60345 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 358 | lthy | 
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 359 | |> Generic_Target.locale_target_const class | 
| 57189 
5140ddfccea7
re-unified approach towards class and locale consts, with refined terminology: foo_const_declaration denotes declaration for a particular logical layer, foo_const the full stack for a particular target
 haftmann parents: 
57188diff
changeset | 360 | (not o (guess_identity orf guess_canonical)) prmode ((b, NoSyn), rhs) | 
| 
5140ddfccea7
re-unified approach towards class and locale consts, with refined terminology: foo_const_declaration denotes declaration for a particular logical layer, foo_const the full stack for a particular target
 haftmann parents: 
57188diff
changeset | 361 | end; | 
| 57072 
dfac6ef0ca28
moved const declaration further down in bootstrap hierarchy: keep Named_Target free of low-level stuff
 haftmann parents: 
57070diff
changeset | 362 | |
| 60345 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 363 | local | 
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 364 | |
| 57119 | 365 | fun dangling_params_for lthy class (type_params, term_params) = | 
| 27690 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 366 | let | 
| 57119 | 367 | val class_param_names = | 
| 368 | map fst (these_params (Proof_Context.theory_of lthy) [class]); | |
| 369 | val dangling_term_params = | |
| 370 | subtract (fn (v, Free (w, _)) => v = w | _ => false) class_param_names term_params; | |
| 57161 
6254c51cd210
formal treatment of dangling parameters for class abbrevs analogously to class consts
 haftmann parents: 
57148diff
changeset | 371 | in (type_params, dangling_term_params) end; | 
| 57119 | 372 | |
| 373 | fun global_def (b, eq) thy = | |
| 374 | thy | |
| 375 | |> Thm.add_def_global false false (b, eq) | |
| 376 | |>> (Thm.varifyT_global o snd) | |
| 377 | |-> (fn def_thm => Global_Theory.store_thm (b, def_thm) | |
| 378 | #> snd | |
| 379 | #> pair def_thm); | |
| 380 | ||
| 57142 | 381 | fun canonical_const class phi dangling_params ((b, mx), rhs) thy = | 
| 57119 | 382 | let | 
| 57147 | 383 | val b_def = Binding.suffix_name "_dict" b; | 
| 384 | val c = Sign.full_name thy b; | |
| 385 | val ty = map Term.fastype_of dangling_params ---> Term.fastype_of rhs; | |
| 386 | val def_eq = Logic.mk_equals (list_comb (Const (c, ty), dangling_params), rhs) | |
| 29577 | 387 | |> map_types Type.strip_sorts; | 
| 27690 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 388 | in | 
| 29577 | 389 | thy | 
| 57147 | 390 | |> Sign.declare_const_global ((b, Type.strip_sorts ty), mx) | 
| 29577 | 391 | |> snd | 
| 57119 | 392 | |> global_def (b_def, def_eq) | 
| 57173 
897cc57a6f55
always refine interpretation morphism using canonical constant's definition theorem
 haftmann parents: 
57161diff
changeset | 393 | |-> (fn def_thm => register_def class def_thm) | 
| 60346 | 394 | |> null dangling_params ? register_operation class (c, rhs) false | 
| 57147 | 395 | |> Sign.add_const_constraint (c, SOME ty) | 
| 27690 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 396 | end; | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 397 | |
| 38619 
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
 haftmann parents: 
38392diff
changeset | 398 | in | 
| 
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
 haftmann parents: 
38392diff
changeset | 399 | |
| 57119 | 400 | fun const class ((b, mx), lhs) params lthy = | 
| 401 | let | |
| 57141 | 402 | val phi = morphism (Proof_Context.theory_of lthy) class; | 
| 57161 
6254c51cd210
formal treatment of dangling parameters for class abbrevs analogously to class consts
 haftmann parents: 
57148diff
changeset | 403 | val dangling_params = map (Morphism.term phi) (uncurry append (dangling_params_for lthy class params)); | 
| 57119 | 404 | in | 
| 405 | lthy | |
| 60337 | 406 | |> target_const class phi Syntax.mode_default (b, lhs) | 
| 57148 
4069c9b3803a
definition in class: provide explicit auxiliary abbreviation carrying potential mixfix syntax in presence of dangling parameters
 haftmann parents: 
57147diff
changeset | 407 | |> Local_Theory.raw_theory (canonical_const class phi dangling_params | 
| 
4069c9b3803a
definition in class: provide explicit auxiliary abbreviation carrying potential mixfix syntax in presence of dangling parameters
 haftmann parents: 
57147diff
changeset | 408 | ((Morphism.binding phi b, if null dangling_params then mx else NoSyn), Morphism.term phi lhs)) | 
| 57192 
180e955711cf
yet another attempt for terminology: foo_target_bar denotes an operation bar operating solely on the target context of target foo, foo_bar denotes a whole stack of operations to accomplish bar for target foo
 haftmann parents: 
57189diff
changeset | 409 | |> Generic_Target.standard_const (fn (this, other) => other <> 0 andalso this <> other) | 
| 57148 
4069c9b3803a
definition in class: provide explicit auxiliary abbreviation carrying potential mixfix syntax in presence of dangling parameters
 haftmann parents: 
57147diff
changeset | 410 | Syntax.mode_default ((b, if null dangling_params then NoSyn else mx), lhs) | 
| 57141 | 411 | |> synchronize_class_syntax_target class | 
| 57119 | 412 | end; | 
| 57072 
dfac6ef0ca28
moved const declaration further down in bootstrap hierarchy: keep Named_Target free of low-level stuff
 haftmann parents: 
57070diff
changeset | 413 | |
| 60345 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 414 | end; | 
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 415 | |
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 416 | local | 
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 417 | |
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 418 | fun canonical_abbrev class phi prmode with_syntax ((b, mx), rhs) thy = | 
| 57141 | 419 | let | 
| 60345 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 420 | val c = Sign.full_name thy b; | 
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 421 | val constrain = map_atyps (fn T as TFree (v, _) => | 
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 422 | if v = Name.aT then TFree (v, [class]) else T | T => T); | 
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 423 | val rhs' = map_types constrain rhs; | 
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 424 | in | 
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 425 | thy | 
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 426 | |> Sign.add_abbrev (#1 prmode) (b, Logic.varify_types_global rhs') | 
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 427 | |> snd | 
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 428 | |> with_syntax ? Sign.notation true prmode [(Const (c, fastype_of rhs), mx)] | 
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 429 | |> with_syntax ? register_operation class (c, rhs) | 
| 60346 | 430 | (#1 prmode = Print_Mode.input) | 
| 60345 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 431 | |> Sign.add_const_constraint (c, SOME (fastype_of rhs')) | 
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 432 | end; | 
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 433 | |
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 434 | fun canonical_abbrev_target class phi prmode ((b, mx), rhs) lthy = | 
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 435 | let | 
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 436 | val thy = Proof_Context.theory_of lthy; | 
| 60347 | 437 | val preprocess = perhaps (try (Pattern.rewrite_term_top thy (these_unchecks thy [class]) [])); | 
| 60345 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 438 | val (global_rhs, (extra_tfrees, (type_params, term_params))) = | 
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 439 | Generic_Target.export_abbrev lthy preprocess rhs; | 
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 440 | val mx' = Generic_Target.check_mixfix_global (b, null term_params) mx; | 
| 57141 | 441 | in | 
| 442 | lthy | |
| 60345 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 443 | |> Local_Theory.raw_theory (canonical_abbrev class phi prmode (null term_params) | 
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 444 | ((Morphism.binding phi b, mx'), Logic.unvarify_types_global global_rhs)) | 
| 57141 | 445 | end; | 
| 38619 
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
 haftmann parents: 
38392diff
changeset | 446 | |
| 60345 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 447 | fun further_abbrev_target class phi prmode (b, mx) rhs params = | 
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 448 | Generic_Target.background_abbrev (b, rhs) (snd params) | 
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 449 | #-> (fn (lhs, _) => target_const class phi prmode (b, lhs) | 
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 450 | #> Generic_Target.standard_const (fn (this, other) => other <> 0 andalso this <> other) prmode ((b, mx), lhs)) | 
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 451 | |
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 452 | in | 
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 453 | |
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 454 | fun abbrev class prmode ((b, mx), rhs) lthy = | 
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 455 | let | 
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 456 | val thy = Proof_Context.theory_of lthy; | 
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 457 | val phi = morphism thy class; | 
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 458 | val rhs_generic = | 
| 60347 | 459 | perhaps (try (Pattern.rewrite_term_top thy (these_unchecks_reversed thy [class]) [])) rhs; | 
| 60345 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 460 | in | 
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 461 | lthy | 
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 462 | |> canonical_abbrev_target class phi prmode ((b, mx), rhs) | 
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 463 | |> Generic_Target.abbrev (further_abbrev_target class phi) prmode ((b, mx), rhs_generic) | 
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 464 | ||> synchronize_class_syntax_target class | 
| 
592806806494
completely separated canonical class abbreviations from abbreviations stemming from non-canonical morphisms -- these have no shared concept
 haftmann parents: 
60344diff
changeset | 465 | end; | 
| 60344 
a40369ea3ba5
self-contained formulation of abbrev for named targets
 haftmann parents: 
60341diff
changeset | 466 | |
| 38619 
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
 haftmann parents: 
38392diff
changeset | 467 | end; | 
| 
25e401d53900
tuned: less formal noise in Named_Target, more coherence in Class
 haftmann parents: 
38392diff
changeset | 468 | |
| 27690 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 469 | |
| 54866 
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
 haftmann parents: 
54742diff
changeset | 470 | (* subclasses *) | 
| 
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
 haftmann parents: 
54742diff
changeset | 471 | |
| 
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
 haftmann parents: 
54742diff
changeset | 472 | fun register_subclass (sub, sup) some_dep_morph some_witn export lthy = | 
| 
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
 haftmann parents: 
54742diff
changeset | 473 | let | 
| 
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
 haftmann parents: 
54742diff
changeset | 474 | val thy = Proof_Context.theory_of lthy; | 
| 
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
 haftmann parents: 
54742diff
changeset | 475 | val intros = (snd o rules thy) sup :: map_filter I | 
| 54883 
dd04a8b654fc
proper context for norm_hhf and derived operations;
 wenzelm parents: 
54866diff
changeset | 476 | [Option.map (Drule.export_without_context_open o Element.conclude_witness lthy) some_witn, | 
| 54866 
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
 haftmann parents: 
54742diff
changeset | 477 | (fst o rules thy) sub]; | 
| 
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
 haftmann parents: 
54742diff
changeset | 478 | val classrel = | 
| 
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
 haftmann parents: 
54742diff
changeset | 479 | Goal.prove_sorry_global thy [] [] (Logic.mk_classrel (sub, sup)) | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59423diff
changeset | 480 |         (fn {context = ctxt, ...} => EVERY (map (TRYALL o resolve_tac ctxt o single) intros));
 | 
| 54866 
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
 haftmann parents: 
54742diff
changeset | 481 | val diff_sort = Sign.complete_sort thy [sup] | 
| 
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
 haftmann parents: 
54742diff
changeset | 482 | |> subtract (op =) (Sign.complete_sort thy [sub]) | 
| 
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
 haftmann parents: 
54742diff
changeset | 483 | |> filter (is_class thy); | 
| 
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
 haftmann parents: 
54742diff
changeset | 484 | fun add_dependency some_wit = case some_dep_morph of | 
| 
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
 haftmann parents: 
54742diff
changeset | 485 | SOME dep_morph => Generic_Target.locale_dependency sub | 
| 
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
 haftmann parents: 
54742diff
changeset | 486 | (sup, dep_morph $> Element.satisfy_morphism (the_list some_witn)) NONE export | 
| 
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
 haftmann parents: 
54742diff
changeset | 487 | | NONE => I; | 
| 
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
 haftmann parents: 
54742diff
changeset | 488 | in | 
| 
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
 haftmann parents: 
54742diff
changeset | 489 | lthy | 
| 
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
 haftmann parents: 
54742diff
changeset | 490 | |> Local_Theory.raw_theory | 
| 
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
 haftmann parents: 
54742diff
changeset | 491 | (Axclass.add_classrel classrel | 
| 
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
 haftmann parents: 
54742diff
changeset | 492 | #> Class_Data.map (Graph.add_edge (sub, sup)) | 
| 
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
 haftmann parents: 
54742diff
changeset | 493 | #> activate_defs sub (these_defs thy diff_sort)) | 
| 
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
 haftmann parents: 
54742diff
changeset | 494 | |> add_dependency some_witn | 
| 
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
 haftmann parents: 
54742diff
changeset | 495 | |> synchronize_class_syntax_target sub | 
| 
7b9a67cbd48f
self-contained formulation of subclass command, avoiding hard-wired Named_Target.init
 haftmann parents: 
54742diff
changeset | 496 | end; | 
| 31635 | 497 | |
| 498 | local | |
| 499 | ||
| 500 | fun gen_classrel mk_prop classrel thy = | |
| 501 | let | |
| 502 | fun after_qed results = | |
| 51685 
385ef6706252
more standard module name Axclass (according to file name);
 wenzelm parents: 
51578diff
changeset | 503 | Proof_Context.background_theory ((fold o fold) Axclass.add_classrel results); | 
| 31635 | 504 | in | 
| 505 | thy | |
| 42360 | 506 | |> Proof_Context.init_global | 
| 36323 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36106diff
changeset | 507 | |> Proof.theorem NONE after_qed [[(mk_prop thy classrel, [])]] | 
| 31635 | 508 | end; | 
| 509 | ||
| 510 | in | |
| 511 | ||
| 512 | val classrel = | |
| 51685 
385ef6706252
more standard module name Axclass (according to file name);
 wenzelm parents: 
51578diff
changeset | 513 | gen_classrel (Logic.mk_classrel oo Axclass.cert_classrel); | 
| 31635 | 514 | val classrel_cmd = | 
| 51685 
385ef6706252
more standard module name Axclass (according to file name);
 wenzelm parents: 
51578diff
changeset | 515 | gen_classrel (Logic.mk_classrel oo Axclass.read_classrel); | 
| 31635 | 516 | |
| 517 | end; (*local*) | |
| 518 | ||
| 519 | ||
| 25462 | 520 | (** instantiation target **) | 
| 521 | ||
| 522 | (* bookkeeping *) | |
| 523 | ||
| 524 | datatype instantiation = Instantiation of {
 | |
| 25864 | 525 | arities: string list * (string * sort) list * sort, | 
| 25462 | 526 | params: ((string * string) * (string * typ)) list | 
| 25603 | 527 | (*(instantiation parameter, type constructor), (local instantiation parameter, typ)*) | 
| 25462 | 528 | } | 
| 529 | ||
| 59150 | 530 | fun make_instantiation (arities, params) = | 
| 531 |   Instantiation {arities = arities, params = params};
 | |
| 532 | ||
| 533 | val empty_instantiation = make_instantiation (([], [], []), []); | |
| 534 | ||
| 33519 | 535 | structure Instantiation = Proof_Data | 
| 25462 | 536 | ( | 
| 46919 | 537 | type T = instantiation; | 
| 59150 | 538 | fun init _ = empty_instantiation; | 
| 25462 | 539 | ); | 
| 540 | ||
| 46919 | 541 | val get_instantiation = | 
| 542 | (fn Instantiation data => data) o Instantiation.get o Local_Theory.target_of; | |
| 25462 | 543 | |
| 46919 | 544 | fun map_instantiation f = | 
| 545 | (Local_Theory.target o Instantiation.map) | |
| 59150 | 546 |     (fn Instantiation {arities, params} => make_instantiation (f (arities, params)));
 | 
| 46919 | 547 | |
| 548 | fun the_instantiation lthy = | |
| 549 | (case get_instantiation lthy of | |
| 550 |     {arities = ([], [], []), ...} => error "No instantiation target"
 | |
| 551 | | data => data); | |
| 25462 | 552 | |
| 25485 | 553 | val instantiation_params = #params o get_instantiation; | 
| 25462 | 554 | |
| 30519 
c05c0199826f
coherent binding policy with primitive target operations
 haftmann parents: 
30364diff
changeset | 555 | fun instantiation_param lthy b = instantiation_params lthy | 
| 
c05c0199826f
coherent binding policy with primitive target operations
 haftmann parents: 
30364diff
changeset | 556 | |> find_first (fn (_, (v, _)) => Binding.name_of b = v) | 
| 25462 | 557 | |> Option.map (fst o fst); | 
| 558 | ||
| 31869 | 559 | fun read_multi_arity thy (raw_tycos, raw_sorts, raw_sort) = | 
| 560 | let | |
| 42360 | 561 | val ctxt = Proof_Context.init_global thy; | 
| 562 | val all_arities = map (fn raw_tyco => Proof_Context.read_arity ctxt | |
| 31869 | 563 | (raw_tyco, raw_sorts, raw_sort)) raw_tycos; | 
| 564 | val tycos = map #1 all_arities; | |
| 565 | val (_, sorts, sort) = hd all_arities; | |
| 43329 
84472e198515
tuned signature: Name.invent and Name.invent_names;
 wenzelm parents: 
42402diff
changeset | 566 | val vs = Name.invent_names Name.context Name.aT sorts; | 
| 31869 | 567 | in (tycos, vs, sort) end; | 
| 568 | ||
| 25462 | 569 | |
| 570 | (* syntax *) | |
| 571 | ||
| 26238 | 572 | fun synchronize_inst_syntax ctxt = | 
| 25462 | 573 | let | 
| 46919 | 574 |     val Instantiation {params, ...} = Instantiation.get ctxt;
 | 
| 31249 | 575 | |
| 51685 
385ef6706252
more standard module name Axclass (according to file name);
 wenzelm parents: 
51578diff
changeset | 576 | val lookup_inst_param = Axclass.lookup_inst_param | 
| 42360 | 577 | (Sign.consts_of (Proof_Context.theory_of ctxt)) params; | 
| 46919 | 578 | fun subst (c, ty) = | 
| 579 | (case lookup_inst_param (c, ty) of | |
| 580 | SOME (v_ty as (_, ty)) => SOME (ty, Free v_ty) | |
| 581 | | NONE => NONE); | |
| 26238 | 582 | val unchecks = | 
| 583 | map (fn ((c, _), v_ty as (_, ty)) => (Free v_ty, Const (c, ty))) params; | |
| 584 | in | |
| 585 | ctxt | |
| 60338 | 586 |     |> Overloading.map_improvable_syntax (fn {primary_constraints, improve, ...} =>
 | 
| 587 |       {primary_constraints = primary_constraints, secondary_constraints = [],
 | |
| 588 | improve = improve, subst = subst, no_subst_in_abbrev_mode = false, | |
| 589 | unchecks = unchecks}) | |
| 26238 | 590 | end; | 
| 25462 | 591 | |
| 42385 
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
 wenzelm parents: 
42383diff
changeset | 592 | fun resort_terms ctxt algebra consts constraints ts = | 
| 38382 | 593 | let | 
| 42385 
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
 wenzelm parents: 
42383diff
changeset | 594 | fun matchings (Const (c_ty as (c, _))) = | 
| 
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
 wenzelm parents: 
42383diff
changeset | 595 | (case constraints c of | 
| 
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
 wenzelm parents: 
42383diff
changeset | 596 | NONE => I | 
| 
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
 wenzelm parents: 
42383diff
changeset | 597 | | SOME sorts => | 
| 
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
 wenzelm parents: 
42383diff
changeset | 598 | fold2 (curry (Sorts.meet_sort algebra)) (Consts.typargs consts c_ty) sorts) | 
| 
b46b47775cbe
simplified Sorts.class_error: plain Proof.context;
 wenzelm parents: 
42383diff
changeset | 599 | | matchings _ = I; | 
| 38382 | 600 | 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 | 601 | handle Sorts.CLASS_ERROR e => error (Sorts.class_error (Context.pretty ctxt) e); | 
| 38382 | 602 | val inst = map_type_tvar | 
| 603 | (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 | 604 | in if Vartab.is_empty tvartab then ts else (map o map_types) inst ts end; | 
| 38382 | 605 | |
| 25462 | 606 | |
| 607 | (* target *) | |
| 608 | ||
| 47078 | 609 | fun define_overloaded (c, U) v (b_def, rhs) = | 
| 51685 
385ef6706252
more standard module name Axclass (according to file name);
 wenzelm parents: 
51578diff
changeset | 610 | Local_Theory.background_theory_result (Axclass.declare_overloaded (c, U) | 
| 
385ef6706252
more standard module name Axclass (according to file name);
 wenzelm parents: 
51578diff
changeset | 611 | ##>> Axclass.define_overloaded b_def (c, rhs)) | 
| 38382 | 612 | ##> (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 | 613 | ##> Local_Theory.map_contexts (K synchronize_inst_syntax); | 
| 38382 | 614 | |
| 47289 | 615 | fun foundation (((b, U), mx), (b_def, rhs)) params lthy = | 
| 46919 | 616 | (case instantiation_param lthy b of | 
| 617 | SOME c => | |
| 618 |       if mx <> NoSyn then error ("Illegal mixfix syntax for overloaded constant " ^ quote c)
 | |
| 619 | else lthy |> define_overloaded (c, U) (Binding.name_of b) (b_def, rhs) | |
| 60341 
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
 haftmann parents: 
60338diff
changeset | 620 | | NONE => lthy |> Generic_Target.theory_target_foundation (((b, U), mx), (b_def, rhs)) params); | 
| 38382 | 621 | |
| 622 | fun pretty lthy = | |
| 26518 | 623 | let | 
| 46919 | 624 |     val {arities = (tycos, vs, sort), params} = the_instantiation lthy;
 | 
| 38382 | 625 | fun pr_arity tyco = Syntax.pretty_arity lthy (tyco, map snd vs, sort); | 
| 626 | fun pr_param ((c, _), (v, ty)) = | |
| 42359 | 627 | Pretty.block (Pretty.breaks | 
| 55304 | 628 | [Pretty.str v, Pretty.str "==", Proof_Context.pretty_const lthy c, | 
| 42359 | 629 | Pretty.str "::", Syntax.pretty_typ lthy ty]); | 
| 60246 | 630 | in | 
| 631 | [Pretty.block | |
| 632 | (Pretty.fbreaks (Pretty.keyword1 "instantiation" :: map pr_arity tycos @ map pr_param params))] | |
| 633 | end; | |
| 26518 | 634 | |
| 38382 | 635 | fun conclude lthy = | 
| 636 | let | |
| 42359 | 637 | val (tycos, vs, sort) = #arities (the_instantiation lthy); | 
| 42360 | 638 | val thy = Proof_Context.theory_of lthy; | 
| 42359 | 639 | val _ = tycos |> List.app (fn tyco => | 
| 46919 | 640 | if Sign.of_sort thy (Type (tyco, map TFree vs), sort) then () | 
| 55304 | 641 |       else error ("Missing instance proof for type " ^ quote (Proof_Context.markup_type lthy tyco)));
 | 
| 38382 | 642 | in lthy end; | 
| 643 | ||
| 644 | fun instantiation (tycos, vs, sort) thy = | |
| 25462 | 645 | let | 
| 25536 | 646 | val _ = if null tycos then error "At least one arity must be given" else (); | 
| 51685 
385ef6706252
more standard module name Axclass (according to file name);
 wenzelm parents: 
51578diff
changeset | 647 | 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 | 648 | fun get_param tyco (param, (_, (c, ty))) = | 
| 51685 
385ef6706252
more standard module name Axclass (according to file name);
 wenzelm parents: 
51578diff
changeset | 649 | if can (Axclass.param_of_inst thy) (c, tyco) | 
| 25603 | 650 | then NONE else SOME ((c, tyco), | 
| 46917 | 651 | (param ^ "_" ^ Long_Name.base_name tyco, map_atyps (K (Type (tyco, map TFree vs))) ty)); | 
| 31249 | 652 | val params = map_product get_param tycos class_params |> map_filter I; | 
| 51578 | 653 | val _ = if null params andalso forall (fn tyco => can (Sign.arity_sorts thy tyco) sort) tycos | 
| 654 | then error "No parameters and no pending instance proof obligations in instantiation." | |
| 655 | else (); | |
| 26518 | 656 | val primary_constraints = map (apsnd | 
| 31249 | 657 | (map_atyps (K (TVar ((Name.aT, 0), [])))) o snd o snd) class_params; | 
| 26518 | 658 | 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 | 659 | |> fold (fn tyco => Sorts.add_arities (Context.pretty_global thy) | 
| 26518 | 660 | (tyco, map (fn class => (class, map snd vs)) sort)) tycos; | 
| 661 | val consts = Sign.consts_of thy; | |
| 662 | val improve_constraints = AList.lookup (op =) | |
| 31249 | 663 | (map (fn (_, (class, (c, _))) => (c, [[class]])) class_params); | 
| 42402 
c7139609b67d
simplified check/uncheck interfaces: result comparison is hardwired by default;
 wenzelm parents: 
42388diff
changeset | 664 | fun resort_check ctxt ts = resort_terms ctxt algebra consts improve_constraints ts; | 
| 51685 
385ef6706252
more standard module name Axclass (according to file name);
 wenzelm parents: 
51578diff
changeset | 665 | val lookup_inst_param = Axclass.lookup_inst_param consts params; | 
| 42388 | 666 | fun improve (c, ty) = | 
| 667 | (case lookup_inst_param (c, ty) of | |
| 668 | SOME (_, ty') => if Sign.typ_instance thy (ty', ty) then SOME (ty, ty') else NONE | |
| 669 | | NONE => NONE); | |
| 25485 | 670 | in | 
| 671 | thy | |
| 56056 
4d46d53566e6
more efficient local theory operations, by imposing a linear change discipline on the main types/consts tables, in order to speed-up Proof_Context.transfer_syntax required for Local_Theory.raw_theory_result;
 wenzelm parents: 
55763diff
changeset | 672 | |> Sign.change_begin | 
| 42360 | 673 | |> Proof_Context.init_global | 
| 59150 | 674 | |> Instantiation.put (make_instantiation ((tycos, vs, sort), params)) | 
| 27281 | 675 | |> fold (Variable.declare_typ o TFree) vs | 
| 31249 | 676 | |> fold (Variable.declare_names o Free o snd) params | 
| 60338 | 677 |     |> (Overloading.map_improvable_syntax) (K {primary_constraints = primary_constraints,
 | 
| 678 | secondary_constraints = [], improve = improve, subst = K NONE, | |
| 679 | no_subst_in_abbrev_mode = false, unchecks = []}) | |
| 39378 
df86b1b4ce10
more precise name for activation of improveable syntax
 haftmann parents: 
39134diff
changeset | 680 | |> Overloading.activate_improvable_syntax | 
| 45429 | 681 | |> Context.proof_map (Syntax_Phases.term_check 0 "resorting" resort_check) | 
| 26238 | 682 | |> synchronize_inst_syntax | 
| 59876 | 683 | |> Local_Theory.init (Sign.naming_of thy) | 
| 38382 | 684 |        {define = Generic_Target.define foundation,
 | 
| 60341 
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
 haftmann parents: 
60338diff
changeset | 685 | notes = Generic_Target.notes Generic_Target.theory_target_notes, | 
| 
fcbd7f0c52c3
clearly separated target primitives (target_foo) from self-contained target operations (foo)
 haftmann parents: 
60338diff
changeset | 686 | abbrev = Generic_Target.abbrev Generic_Target.theory_target_abbrev, | 
| 47279 
4bab649dedf0
clarified standard_declaration vs. theory_declaration;
 wenzelm parents: 
47250diff
changeset | 687 | declaration = K Generic_Target.theory_declaration, | 
| 56723 
a8f71445c265
subscription as target-specific implementation device
 haftmann parents: 
56334diff
changeset | 688 | subscription = Generic_Target.theory_registration, | 
| 38382 | 689 | pretty = pretty, | 
| 56056 
4d46d53566e6
more efficient local theory operations, by imposing a linear change discipline on the main types/consts tables, in order to speed-up Proof_Context.transfer_syntax required for Local_Theory.raw_theory_result;
 wenzelm parents: 
55763diff
changeset | 690 | exit = conclude #> Local_Theory.target_of #> Sign.change_end_local} | 
| 25485 | 691 | end; | 
| 692 | ||
| 38382 | 693 | fun instantiation_cmd arities thy = | 
| 694 | instantiation (read_multi_arity thy arities) thy; | |
| 26238 | 695 | |
| 25485 | 696 | fun gen_instantiation_instance do_proof after_qed lthy = | 
| 697 | let | |
| 25864 | 698 | val (tycos, vs, sort) = (#arities o the_instantiation) lthy; | 
| 699 | val arities_proof = maps (fn tyco => Logic.mk_arities (tyco, map snd vs, sort)) tycos; | |
| 25462 | 700 | fun after_qed' results = | 
| 51685 
385ef6706252
more standard module name Axclass (according to file name);
 wenzelm parents: 
51578diff
changeset | 701 | Local_Theory.background_theory (fold (Axclass.add_arity o Thm.varifyT_global) results) | 
| 25462 | 702 | #> after_qed; | 
| 703 | in | |
| 704 | lthy | |
| 705 | |> do_proof after_qed' arities_proof | |
| 706 | end; | |
| 707 | ||
| 25485 | 708 | 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 | 709 | Proof.theorem NONE (after_qed o map the_single) (map (fn t => [(t, [])]) ts)); | 
| 25462 | 710 | |
| 25485 | 711 | fun prove_instantiation_instance tac = gen_instantiation_instance (fn after_qed => | 
| 25502 | 712 | fn ts => fn lthy => after_qed (map (fn t => Goal.prove lthy [] [] t | 
| 713 |     (fn {context, ...} => tac context)) ts) lthy) I;
 | |
| 25462 | 714 | |
| 28666 | 715 | fun prove_instantiation_exit tac = prove_instantiation_instance tac | 
| 33671 | 716 | #> Local_Theory.exit_global; | 
| 28666 | 717 | |
| 718 | fun prove_instantiation_exit_result f tac x lthy = | |
| 719 | let | |
| 42360 | 720 | val morph = Proof_Context.export_morphism lthy | 
| 721 | (Proof_Context.init_global (Proof_Context.theory_of lthy)); | |
| 29439 | 722 | val y = f morph x; | 
| 28666 | 723 | in | 
| 724 | lthy | |
| 725 | |> prove_instantiation_exit (fn ctxt => tac ctxt y) | |
| 726 | |> pair y | |
| 727 | end; | |
| 728 | ||
| 29526 | 729 | |
| 31635 | 730 | (* simplified instantiation interface with no class parameter *) | 
| 731 | ||
| 31869 | 732 | fun instance_arity_cmd raw_arities thy = | 
| 31635 | 733 | let | 
| 31869 | 734 | val (tycos, vs, sort) = read_multi_arity thy raw_arities; | 
| 735 | val sorts = map snd vs; | |
| 736 | val arities = maps (fn tyco => Logic.mk_arities (tyco, sorts, sort)) tycos; | |
| 47078 | 737 | fun after_qed results = | 
| 51685 
385ef6706252
more standard module name Axclass (according to file name);
 wenzelm parents: 
51578diff
changeset | 738 | Proof_Context.background_theory ((fold o fold) Axclass.add_arity results); | 
| 31635 | 739 | in | 
| 740 | thy | |
| 42360 | 741 | |> Proof_Context.init_global | 
| 36323 
655e2d74de3a
modernized naming conventions of main Isar proof elements;
 wenzelm parents: 
36106diff
changeset | 742 | |> Proof.theorem NONE after_qed (map (fn t => [(t, [])]) arities) | 
| 31635 | 743 | end; | 
| 744 | ||
| 745 | ||
| 60618 
4c79543cc376
renamed "default" to "standard", to make semantically clear what it is;
 wenzelm parents: 
60347diff
changeset | 746 | |
| 29526 | 747 | (** tactics and methods **) | 
| 748 | ||
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59423diff
changeset | 749 | fun intro_classes_tac ctxt facts st = | 
| 29526 | 750 | let | 
| 751 | val thy = Thm.theory_of_thm st; | |
| 752 | val classes = Sign.all_classes thy; | |
| 753 | val class_trivs = map (Thm.class_triv thy) classes; | |
| 51685 
385ef6706252
more standard module name Axclass (according to file name);
 wenzelm parents: 
51578diff
changeset | 754 | val class_intros = map_filter (try (#intro o Axclass.get_info thy)) classes; | 
| 29526 | 755 | val assm_intros = all_assm_intros thy; | 
| 756 | in | |
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59423diff
changeset | 757 | Method.intros_tac ctxt (class_trivs @ class_intros @ assm_intros) facts st | 
| 29526 | 758 | end; | 
| 759 | ||
| 60618 
4c79543cc376
renamed "default" to "standard", to make semantically clear what it is;
 wenzelm parents: 
60347diff
changeset | 760 | fun standard_intro_classes_tac ctxt facts st = | 
| 
4c79543cc376
renamed "default" to "standard", to make semantically clear what it is;
 wenzelm parents: 
60347diff
changeset | 761 | if null facts andalso not (Thm.no_prems st) then | 
| 
4c79543cc376
renamed "default" to "standard", to make semantically clear what it is;
 wenzelm parents: 
60347diff
changeset | 762 | (intro_classes_tac ctxt [] ORELSE Locale.intro_locales_tac true ctxt []) st | 
| 
4c79543cc376
renamed "default" to "standard", to make semantically clear what it is;
 wenzelm parents: 
60347diff
changeset | 763 | else no_tac st; | 
| 29526 | 764 | |
| 60619 
7512716f03db
no arguments for "standard" (or old "default") methods;
 wenzelm parents: 
60618diff
changeset | 765 | fun standard_tac ctxt facts = | 
| 
7512716f03db
no arguments for "standard" (or old "default") methods;
 wenzelm parents: 
60618diff
changeset | 766 | HEADGOAL (Method.some_rule_tac ctxt [] facts) ORELSE | 
| 60618 
4c79543cc376
renamed "default" to "standard", to make semantically clear what it is;
 wenzelm parents: 
60347diff
changeset | 767 | standard_intro_classes_tac ctxt facts; | 
| 
4c79543cc376
renamed "default" to "standard", to make semantically clear what it is;
 wenzelm parents: 
60347diff
changeset | 768 | |
| 60619 
7512716f03db
no arguments for "standard" (or old "default") methods;
 wenzelm parents: 
60618diff
changeset | 769 | fun default_tac ctxt facts st = | 
| 60618 
4c79543cc376
renamed "default" to "standard", to make semantically clear what it is;
 wenzelm parents: 
60347diff
changeset | 770 | (if Method.detect_closure_state st then | 
| 
4c79543cc376
renamed "default" to "standard", to make semantically clear what it is;
 wenzelm parents: 
60347diff
changeset | 771 | legacy_feature | 
| 
4c79543cc376
renamed "default" to "standard", to make semantically clear what it is;
 wenzelm parents: 
60347diff
changeset | 772 |       ("Old method \"default\"; use \"standard\" instead" ^ Position.here (Position.thread_data ()))
 | 
| 
4c79543cc376
renamed "default" to "standard", to make semantically clear what it is;
 wenzelm parents: 
60347diff
changeset | 773 | else (); | 
| 60619 
7512716f03db
no arguments for "standard" (or old "default") methods;
 wenzelm parents: 
60618diff
changeset | 774 | standard_tac ctxt facts st); | 
| 29526 | 775 | |
| 53171 | 776 | val _ = Theory.setup | 
| 59498 
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
 wenzelm parents: 
59423diff
changeset | 777 |  (Method.setup @{binding intro_classes} (Scan.succeed (METHOD o intro_classes_tac))
 | 
| 30515 | 778 | "back-chain introduction rules of classes" #> | 
| 60619 
7512716f03db
no arguments for "standard" (or old "default") methods;
 wenzelm parents: 
60618diff
changeset | 779 |   Method.setup @{binding standard} (Scan.succeed (METHOD o standard_tac))
 | 
| 60618 
4c79543cc376
renamed "default" to "standard", to make semantically clear what it is;
 wenzelm parents: 
60347diff
changeset | 780 | "standard proof step: Pure intro/elim rule or class introduction" #> | 
| 60619 
7512716f03db
no arguments for "standard" (or old "default") methods;
 wenzelm parents: 
60618diff
changeset | 781 |   Method.setup @{binding default} (Scan.succeed (METHOD o default_tac))
 | 
| 60618 
4c79543cc376
renamed "default" to "standard", to make semantically clear what it is;
 wenzelm parents: 
60347diff
changeset | 782 | "standard proof step: Pure intro/elim rule or class introduction (legacy)"); | 
| 29526 | 783 | |
| 59296 
002d817b4c37
formal pretty bodies for class specifications, accepting additional formal bookkeeping in locale.ML
 haftmann parents: 
59150diff
changeset | 784 | |
| 60249 
09377954133b
tuned output to resemble input syntax more closely;
 wenzelm parents: 
60246diff
changeset | 785 | |
| 59296 
002d817b4c37
formal pretty bodies for class specifications, accepting additional formal bookkeeping in locale.ML
 haftmann parents: 
59150diff
changeset | 786 | (** diagnostics **) | 
| 
002d817b4c37
formal pretty bodies for class specifications, accepting additional formal bookkeeping in locale.ML
 haftmann parents: 
59150diff
changeset | 787 | |
| 60249 
09377954133b
tuned output to resemble input syntax more closely;
 wenzelm parents: 
60246diff
changeset | 788 | fun pretty_specification thy class = | 
| 
09377954133b
tuned output to resemble input syntax more closely;
 wenzelm parents: 
60246diff
changeset | 789 | if is_class thy class then | 
| 59383 
1434ef1e0ede
clarified Class.pretty_specification: imitate input source;
 wenzelm parents: 
59296diff
changeset | 790 | let | 
| 60249 
09377954133b
tuned output to resemble input syntax more closely;
 wenzelm parents: 
60246diff
changeset | 791 | val class_ctxt = init class thy; | 
| 
09377954133b
tuned output to resemble input syntax more closely;
 wenzelm parents: 
60246diff
changeset | 792 | val prt_class = Name_Space.pretty class_ctxt (Proof_Context.class_space class_ctxt); | 
| 
09377954133b
tuned output to resemble input syntax more closely;
 wenzelm parents: 
60246diff
changeset | 793 | |
| 
09377954133b
tuned output to resemble input syntax more closely;
 wenzelm parents: 
60246diff
changeset | 794 | val super_classes = Sign.minimize_sort thy (Sign.super_classes thy class); | 
| 59383 
1434ef1e0ede
clarified Class.pretty_specification: imitate input source;
 wenzelm parents: 
59296diff
changeset | 795 | |
| 
1434ef1e0ede
clarified Class.pretty_specification: imitate input source;
 wenzelm parents: 
59296diff
changeset | 796 | val fix_args = | 
| 60249 
09377954133b
tuned output to resemble input syntax more closely;
 wenzelm parents: 
60246diff
changeset | 797 | #params (Axclass.get_info thy class) | 
| 59383 
1434ef1e0ede
clarified Class.pretty_specification: imitate input source;
 wenzelm parents: 
59296diff
changeset | 798 | |> map (fn (c, T) => (Binding.name (Long_Name.base_name c), SOME T, NoSyn)); | 
| 
1434ef1e0ede
clarified Class.pretty_specification: imitate input source;
 wenzelm parents: 
59296diff
changeset | 799 | val fixes = if null fix_args then [] else [Element.Fixes fix_args]; | 
| 60249 
09377954133b
tuned output to resemble input syntax more closely;
 wenzelm parents: 
60246diff
changeset | 800 | val assumes = Locale.hyp_spec_of thy class; | 
| 59383 
1434ef1e0ede
clarified Class.pretty_specification: imitate input source;
 wenzelm parents: 
59296diff
changeset | 801 | |
| 
1434ef1e0ede
clarified Class.pretty_specification: imitate input source;
 wenzelm parents: 
59296diff
changeset | 802 | val header = | 
| 60249 
09377954133b
tuned output to resemble input syntax more closely;
 wenzelm parents: 
60246diff
changeset | 803 | [Pretty.keyword1 "class", Pretty.brk 1, prt_class class, Pretty.str " =", Pretty.brk 1] @ | 
| 
09377954133b
tuned output to resemble input syntax more closely;
 wenzelm parents: 
60246diff
changeset | 804 | Pretty.separate " +" (map prt_class super_classes) @ | 
| 
09377954133b
tuned output to resemble input syntax more closely;
 wenzelm parents: 
60246diff
changeset | 805 | (if null super_classes then [] else [Pretty.str " +"]); | 
| 59383 
1434ef1e0ede
clarified Class.pretty_specification: imitate input source;
 wenzelm parents: 
59296diff
changeset | 806 | val body = | 
| 59385 | 807 | if null fixes andalso null assumes then [] | 
| 808 | else | |
| 809 | maps (Element.pretty_ctxt_no_attribs class_ctxt) (fixes @ assumes) | |
| 810 | |> maps (fn prt => [Pretty.fbrk, prt]); | |
| 59423 
3a0044e95eba
backed out obsolete workaround from ef1edfb36af7
 haftmann parents: 
59420diff
changeset | 811 | in if null body then [] else [Pretty.block (header @ body)] end | 
| 
3a0044e95eba
backed out obsolete workaround from ef1edfb36af7
 haftmann parents: 
59420diff
changeset | 812 | else []; | 
| 59296 
002d817b4c37
formal pretty bodies for class specifications, accepting additional formal bookkeeping in locale.ML
 haftmann parents: 
59150diff
changeset | 813 | |
| 24218 | 814 | end; |