| author | wenzelm | 
| Sun, 16 Nov 2008 22:12:44 +0100 | |
| changeset 28817 | c8cc94a470d4 | 
| parent 28739 | bbb5f83ce602 | 
| child 28822 | 7ca11ecbc4fb | 
| permissions | -rw-r--r-- | 
| 24218 | 1 | (* Title: Pure/Isar/class.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Florian Haftmann, TU Muenchen | |
| 4 | ||
| 5 | Type classes derived from primitive axclasses and locales. | |
| 6 | *) | |
| 7 | ||
| 8 | signature CLASS = | |
| 9 | sig | |
| 25462 | 10 | (*classes*) | 
| 26247 | 11 | val class: bstring -> class list -> Element.context_i list | 
| 26518 | 12 | -> theory -> string * Proof.context | 
| 26247 | 13 | val class_cmd: bstring -> xstring list -> Element.context list | 
| 26518 | 14 | -> theory -> string * Proof.context | 
| 25485 | 15 | |
| 25311 | 16 | val init: class -> theory -> Proof.context | 
| 28017 | 17 | val declare: class -> Properties.T | 
| 27690 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 18 | -> (string * mixfix) * term -> theory -> theory | 
| 28017 | 19 | val abbrev: class -> Syntax.mode -> Properties.T | 
| 25104 | 20 | -> (string * mixfix) * term -> theory -> theory | 
| 25083 | 21 | val refresh_syntax: class -> Proof.context -> Proof.context | 
| 25485 | 22 | |
| 24589 | 23 | val intro_classes_tac: thm list -> tactic | 
| 26470 | 24 | val default_intro_tac: Proof.context -> thm list -> tactic | 
| 27684 
f45fd1159a4b
subclass now also works for subclasses with empty specificaton
 haftmann parents: 
27281diff
changeset | 25 | val prove_subclass: class * class -> thm option -> theory -> theory | 
| 25485 | 26 | |
| 27 | val class_prefix: string -> string | |
| 28 | val is_class: theory -> class -> bool | |
| 26518 | 29 | val these_params: theory -> sort -> (string * (class * (string * typ))) list | 
| 24589 | 30 | val print_classes: theory -> unit | 
| 24423 
ae9cd0e92423
overloaded definitions accompanied by explicit constants
 haftmann parents: 
24304diff
changeset | 31 | |
| 25462 | 32 | (*instances*) | 
| 26247 | 33 | val init_instantiation: string list * (string * sort) list * sort | 
| 34 | -> theory -> local_theory | |
| 35 | val instantiation_instance: (local_theory -> local_theory) | |
| 36 | -> local_theory -> Proof.state | |
| 37 | val prove_instantiation_instance: (Proof.context -> tactic) | |
| 38 | -> local_theory -> local_theory | |
| 28666 | 39 | val prove_instantiation_exit: (Proof.context -> tactic) | 
| 40 | -> local_theory -> theory | |
| 41 | val prove_instantiation_exit_result: (morphism -> 'a -> 'b) | |
| 42 | -> (Proof.context -> 'b -> tactic) -> 'a -> local_theory -> 'b * theory | |
| 25485 | 43 | val conclude_instantiation: local_theory -> local_theory | 
| 25603 | 44 | val instantiation_param: local_theory -> string -> string option | 
| 25485 | 45 | val confirm_declaration: string -> local_theory -> local_theory | 
| 25603 | 46 | val pretty_instantiation: local_theory -> Pretty.T | 
| 26259 | 47 | val type_name: string -> string | 
| 25485 | 48 | |
| 25462 | 49 | (*old axclass layer*) | 
| 50 | val axclass_cmd: bstring * xstring list | |
| 28084 
a05ca48ef263
type Attrib.binding abbreviates Name.binding without attributes;
 wenzelm parents: 
28083diff
changeset | 51 | -> (Attrib.binding * string list) list | 
| 25462 | 52 | -> theory -> class * theory | 
| 53 | val classrel_cmd: xstring * xstring -> theory -> Proof.state | |
| 54 | ||
| 55 | (*old instance layer*) | |
| 25536 | 56 | val instance_arity: (theory -> theory) -> arity -> theory -> Proof.state | 
| 57 | val instance_arity_cmd: bstring * xstring list * xstring -> theory -> Proof.state | |
| 24218 | 58 | end; | 
| 59 | ||
| 60 | structure Class : CLASS = | |
| 61 | struct | |
| 62 | ||
| 63 | (** auxiliary **) | |
| 64 | ||
| 25002 | 65 | fun prove_interpretation tac prfx_atts expr inst = | 
| 28715 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 66 | Locale.interpretation_i I prfx_atts expr inst | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 67 | ##> Proof.global_terminal_proof | 
| 27761 
b95e9ba0ca1d
Interpretation command (theory/proof context) no longer simplifies goal.
 ballarin parents: 
27708diff
changeset | 68 | (Method.Basic (fn ctxt => Method.SIMPLE_METHOD (tac ctxt), Position.none), NONE) | 
| 28715 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 69 | ##> ProofContext.theory_of; | 
| 24589 | 70 | |
| 25195 
62638dcafe38
fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
 haftmann parents: 
25163diff
changeset | 71 | fun prove_interpretation_in tac after_qed (name, expr) = | 
| 
62638dcafe38
fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
 haftmann parents: 
25163diff
changeset | 72 | Locale.interpretation_in_locale | 
| 
62638dcafe38
fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
 haftmann parents: 
25163diff
changeset | 73 | (ProofContext.theory after_qed) (name, expr) | 
| 
62638dcafe38
fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
 haftmann parents: 
25163diff
changeset | 74 | #> Proof.global_terminal_proof | 
| 
62638dcafe38
fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
 haftmann parents: 
25163diff
changeset | 75 | (Method.Basic (K (Method.SIMPLE_METHOD tac), Position.none), NONE) | 
| 
62638dcafe38
fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
 haftmann parents: 
25163diff
changeset | 76 | #> ProofContext.theory_of; | 
| 
62638dcafe38
fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
 haftmann parents: 
25163diff
changeset | 77 | |
| 28715 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 78 | val class_prefix = Logic.const_of_class o Sign.base_name; | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 79 | |
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 80 | fun activate_class_morphism thy class inst morphism = | 
| 28739 | 81 | Locale.get_interpret_morph thy class (false, class_prefix class) "" morphism class inst; | 
| 28715 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 82 | |
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 83 | fun prove_class_interpretation class inst consts hyp_facts def_facts thy = | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 84 | let | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 85 | val constraints = map (fn (class, c) => map_atyps (K (TFree (Name.aT, | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 86 | [class]))) (Sign.the_const_type thy c)) consts; | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 87 | val no_constraints = map (map_atyps (K (TFree (Name.aT, [])))) constraints; | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 88 | fun add_constraint c T = Sign.add_const_constraint (c, SOME T); | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 89 | fun tac ctxt = ALLGOALS (ProofContext.fact_tac (hyp_facts @ def_facts) | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 90 | ORELSE' (fn n => SELECT_GOAL (Locale.intro_locales_tac false ctxt []) n)); | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 91 | val prfx = class_prefix class; | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 92 | in | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 93 | thy | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 94 | |> fold2 add_constraint (map snd consts) no_constraints | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 95 | |> prove_interpretation tac (false, prfx) (Locale.Locale class) | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 96 | (map SOME inst, map (pair (Attrib.no_binding) o Thm.prop_of) def_facts) | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 97 | ||> fold2 add_constraint (map snd consts) constraints | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 98 | end; | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 99 | |
| 24589 | 100 | |
| 25485 | 101 | (** primitive axclass and instance commands **) | 
| 24589 | 102 | |
| 24218 | 103 | fun axclass_cmd (class, raw_superclasses) raw_specs thy = | 
| 104 | let | |
| 105 | val ctxt = ProofContext.init thy; | |
| 106 | val superclasses = map (Sign.read_class thy) raw_superclasses; | |
| 24589 | 107 | val name_atts = map ((apsnd o map) (Attrib.attribute thy) o fst) | 
| 108 | raw_specs; | |
| 109 | val axiomss = ProofContext.read_propp (ctxt, map (map (rpair []) o snd) | |
| 110 | raw_specs) | |
| 24218 | 111 | |> snd | 
| 112 | |> (map o map) fst; | |
| 24589 | 113 | in | 
| 114 | AxClass.define_class (class, superclasses) [] | |
| 115 | (name_atts ~~ axiomss) thy | |
| 116 | end; | |
| 24218 | 117 | |
| 118 | local | |
| 119 | ||
| 120 | fun gen_instance mk_prop add_thm after_qed insts thy = | |
| 121 | let | |
| 122 | fun after_qed' results = | |
| 123 | ProofContext.theory ((fold o fold) add_thm results #> after_qed); | |
| 124 | in | |
| 125 | thy | |
| 126 | |> ProofContext.init | |
| 24589 | 127 | |> Proof.theorem_i NONE after_qed' ((map (fn t => [(t, [])]) | 
| 25536 | 128 | o mk_prop thy) insts) | 
| 24218 | 129 | end; | 
| 130 | ||
| 131 | in | |
| 132 | ||
| 24589 | 133 | val instance_arity = | 
| 24218 | 134 | gen_instance (Logic.mk_arities oo Sign.cert_arity) AxClass.add_arity; | 
| 25502 | 135 | val instance_arity_cmd = | 
| 136 | gen_instance (Logic.mk_arities oo Sign.read_arity) AxClass.add_arity I; | |
| 24589 | 137 | val classrel = | 
| 25536 | 138 | gen_instance (single oo (Logic.mk_classrel oo AxClass.cert_classrel)) AxClass.add_classrel I; | 
| 24589 | 139 | val classrel_cmd = | 
| 25536 | 140 | gen_instance (single oo (Logic.mk_classrel oo AxClass.read_classrel)) AxClass.add_classrel I; | 
| 24218 | 141 | |
| 142 | end; (*local*) | |
| 143 | ||
| 144 | ||
| 24589 | 145 | (** class data **) | 
| 24218 | 146 | |
| 147 | datatype class_data = ClassData of {
 | |
| 28715 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 148 | |
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 149 | (* static part *) | 
| 24218 | 150 | consts: (string * string) list | 
| 24836 | 151 | (*locale parameter ~> constant name*), | 
| 25062 | 152 | base_sort: sort, | 
| 28715 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 153 | inst: term list | 
| 25083 | 154 | (*canonical interpretation*), | 
| 28715 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 155 | morphism: Morphism.morphism * ((typ Vartab.table * typ list) * (term Vartab.table * term list)) | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 156 | (*morphism cookie of canonical interpretation*), | 
| 25618 | 157 | assm_intro: thm option, | 
| 158 | of_class: thm, | |
| 159 | axiom: thm option, | |
| 28715 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 160 | |
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 161 | (* dynamic part *) | 
| 24657 | 162 | defs: thm list, | 
| 25368 
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
 haftmann parents: 
25344diff
changeset | 163 | operations: (string * (class * (typ * term))) list | 
| 28715 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 164 | |
| 24657 | 165 | }; | 
| 24218 | 166 | |
| 24657 | 167 | fun rep_class_data (ClassData d) = d; | 
| 25618 | 168 | fun mk_class_data ((consts, base_sort, inst, morphism, assm_intro, of_class, axiom), | 
| 25368 
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
 haftmann parents: 
25344diff
changeset | 169 | (defs, operations)) = | 
| 25062 | 170 |   ClassData { consts = consts, base_sort = base_sort, inst = inst,
 | 
| 26463 | 171 | morphism = morphism, assm_intro = assm_intro, of_class = of_class, axiom = axiom, | 
| 25618 | 172 | defs = defs, operations = operations }; | 
| 28715 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 173 | fun map_class_data f (ClassData { consts, base_sort, inst, morphism, assm_intro,
 | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 174 | of_class, axiom, defs, operations }) = | 
| 25618 | 175 | mk_class_data (f ((consts, base_sort, inst, morphism, assm_intro, of_class, axiom), | 
| 25368 
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
 haftmann parents: 
25344diff
changeset | 176 | (defs, operations))); | 
| 25038 | 177 | fun merge_class_data _ (ClassData { consts = consts,
 | 
| 25618 | 178 | base_sort = base_sort, inst = inst, morphism = morphism, assm_intro = assm_intro, | 
| 179 | of_class = of_class, axiom = axiom, defs = defs1, operations = operations1 }, | |
| 180 |   ClassData { consts = _, base_sort = _, inst = _, morphism = _, assm_intro = _,
 | |
| 181 | of_class = _, axiom = _, defs = defs2, operations = operations2 }) = | |
| 182 | mk_class_data ((consts, base_sort, inst, morphism, assm_intro, of_class, axiom), | |
| 24914 
95cda5dd58d5
added proper subclass concept; improved class target
 haftmann parents: 
24901diff
changeset | 183 | (Thm.merge_thms (defs1, defs2), | 
| 25368 
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
 haftmann parents: 
25344diff
changeset | 184 | AList.merge (op =) (K true) (operations1, operations2))); | 
| 24218 | 185 | |
| 186 | structure ClassData = TheoryDataFun | |
| 187 | ( | |
| 25038 | 188 | type T = class_data Graph.T | 
| 189 | val empty = Graph.empty; | |
| 24218 | 190 | val copy = I; | 
| 191 | val extend = I; | |
| 25038 | 192 | fun merge _ = Graph.join merge_class_data; | 
| 24218 | 193 | ); | 
| 194 | ||
| 195 | ||
| 196 | (* queries *) | |
| 197 | ||
| 25038 | 198 | val lookup_class_data = Option.map rep_class_data oo try o Graph.get_node o ClassData.get; | 
| 24218 | 199 | |
| 24589 | 200 | fun the_class_data thy class = case lookup_class_data thy class | 
| 25020 | 201 |  of NONE => error ("Undeclared class " ^ quote class)
 | 
| 24589 | 202 | | SOME data => data; | 
| 24218 | 203 | |
| 25038 | 204 | val is_class = is_some oo lookup_class_data; | 
| 205 | ||
| 206 | val ancestry = Graph.all_succs o ClassData.get; | |
| 24218 | 207 | |
| 28715 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 208 | fun the_inst thy = #inst o the_class_data thy; | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 209 | |
| 25002 | 210 | fun these_params thy = | 
| 24218 | 211 | let | 
| 212 | fun params class = | |
| 213 | let | |
| 24930 
cc2e0e8c81af
renamed AxClass.get_definition to AxClass.get_info (again);
 wenzelm parents: 
24920diff
changeset | 214 | val const_typs = (#params o AxClass.get_info thy) class; | 
| 24657 | 215 | val const_names = (#consts o the_class_data thy) class; | 
| 24218 | 216 | in | 
| 26518 | 217 | (map o apsnd) | 
| 218 | (fn c => (class, (c, (the o AList.lookup (op =) const_typs) c))) const_names | |
| 24218 | 219 | end; | 
| 220 | in maps params o ancestry thy end; | |
| 221 | ||
| 25618 | 222 | fun these_assm_intros thy = | 
| 223 | Graph.fold (fn (_, (data, _)) => fold (insert Thm.eq_thm) | |
| 224 | ((the_list o #assm_intro o rep_class_data) data)) (ClassData.get thy) []; | |
| 24218 | 225 | |
| 28715 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 226 | fun these_defs thy = maps (these o Option.map #defs o lookup_class_data thy) o ancestry thy; | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 227 | |
| 24836 | 228 | fun these_operations thy = | 
| 229 | maps (#operations o the_class_data thy) o ancestry thy; | |
| 24657 | 230 | |
| 28715 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 231 | fun morphism thy class = | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 232 | let | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 233 |     val { inst, morphism, ... } = the_class_data thy class;
 | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 234 | in activate_class_morphism thy class inst morphism end; | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 235 | |
| 24218 | 236 | fun print_classes thy = | 
| 237 | let | |
| 24920 | 238 | val ctxt = ProofContext.init thy; | 
| 24218 | 239 | val algebra = Sign.classes_of thy; | 
| 240 | val arities = | |
| 241 | Symtab.empty | |
| 242 | |> Symtab.fold (fn (tyco, arities) => fold (fn (class, _) => | |
| 243 | Symtab.map_default (class, []) (insert (op =) tyco)) arities) | |
| 244 | ((#arities o Sorts.rep_algebra) algebra); | |
| 245 | val the_arities = these o Symtab.lookup arities; | |
| 246 | fun mk_arity class tyco = | |
| 247 | let | |
| 248 | val Ss = Sorts.mg_domain algebra tyco [class]; | |
| 24920 | 249 | in Syntax.pretty_arity ctxt (tyco, Ss, [class]) end; | 
| 24218 | 250 | fun mk_param (c, ty) = Pretty.str (Sign.extern_const thy c ^ " :: " | 
| 24920 | 251 | ^ setmp show_sorts false (Syntax.string_of_typ ctxt o Type.strip_sorts) ty); | 
| 24218 | 252 | fun mk_entry class = (Pretty.block o Pretty.fbreaks o map_filter I) [ | 
| 25062 | 253 |       (SOME o Pretty.str) ("class " ^ Sign.extern_class thy class ^ ":"),
 | 
| 24218 | 254 | (SOME o Pretty.block) [Pretty.str "supersort: ", | 
| 24920 | 255 | (Syntax.pretty_sort ctxt o Sign.minimize_sort thy o Sign.super_classes thy) class], | 
| 25062 | 256 | if is_class thy class then (SOME o Pretty.str) | 
| 257 |         ("locale: " ^ Locale.extern thy class) else NONE,
 | |
| 258 | ((fn [] => NONE | ps => (SOME o Pretty.block o Pretty.fbreaks) | |
| 259 | (Pretty.str "parameters:" :: ps)) o map mk_param | |
| 24930 
cc2e0e8c81af
renamed AxClass.get_definition to AxClass.get_info (again);
 wenzelm parents: 
24920diff
changeset | 260 | o these o Option.map #params o try (AxClass.get_info thy)) class, | 
| 24218 | 261 | (SOME o Pretty.block o Pretty.breaks) [ | 
| 262 | Pretty.str "instances:", | |
| 263 | Pretty.list "" "" (map (mk_arity class) (the_arities class)) | |
| 264 | ] | |
| 265 | ] | |
| 266 | in | |
| 24589 | 267 | (Pretty.writeln o Pretty.chunks o separate (Pretty.str "") | 
| 268 | o map mk_entry o Sorts.all_classes) algebra | |
| 24218 | 269 | end; | 
| 270 | ||
| 271 | ||
| 272 | (* updaters *) | |
| 273 | ||
| 25618 | 274 | fun add_class_data ((class, superclasses), | 
| 25711 | 275 | (params, base_sort, inst, phi, axiom, assm_intro, of_class)) thy = | 
| 25002 | 276 | let | 
| 25368 
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
 haftmann parents: 
25344diff
changeset | 277 | val operations = map (fn (v_ty as (_, ty), (c, _)) => | 
| 25683 | 278 | (c, (class, (ty, Free v_ty)))) params; | 
| 25038 | 279 | val add_class = Graph.new_node (class, | 
| 25683 | 280 | mk_class_data (((map o pairself) fst params, base_sort, | 
| 28715 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 281 | inst, phi, assm_intro, of_class, axiom), ([], operations))) | 
| 25002 | 282 | #> fold (curry Graph.add_edge class) superclasses; | 
| 25618 | 283 | in ClassData.map add_class thy end; | 
| 24218 | 284 | |
| 25368 
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
 haftmann parents: 
25344diff
changeset | 285 | fun register_operation class (c, (t, some_def)) thy = | 
| 25062 | 286 | let | 
| 25368 
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
 haftmann parents: 
25344diff
changeset | 287 | val base_sort = (#base_sort o the_class_data thy) class; | 
| 26518 | 288 | val prep_typ = map_type_tvar | 
| 289 | (fn (vi as (v, _), sort) => if Name.aT = v | |
| 25368 
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
 haftmann parents: 
25344diff
changeset | 290 | then TFree (v, base_sort) else TVar (vi, sort)); | 
| 
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
 haftmann parents: 
25344diff
changeset | 291 | val t' = map_types prep_typ t; | 
| 
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
 haftmann parents: 
25344diff
changeset | 292 | val ty' = Term.fastype_of t'; | 
| 25062 | 293 | in | 
| 294 | thy | |
| 295 | |> (ClassData.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 | 296 | (fn (defs, operations) => | 
| 25096 | 297 | (fold cons (the_list some_def) defs, | 
| 25368 
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
 haftmann parents: 
25344diff
changeset | 298 | (c, (class, (ty', t'))) :: operations)) | 
| 25062 | 299 | end; | 
| 24218 | 300 | |
| 24589 | 301 | |
| 302 | (** rule calculation, tactics and methods **) | |
| 303 | ||
| 28715 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 304 | fun calculate_axiom thy sups base_sort assm_axiom param_map class = | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 305 | case Locale.intros thy class | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 306 | of (_, []) => assm_axiom | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 307 | | (_, [intro]) => | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 308 | let | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 309 | fun instantiate thy sort = Thm.instantiate ([pairself (Thm.ctyp_of thy o TVar o pair (Name.aT, 0)) | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 310 | (base_sort, sort)], map (fn (v, (c, ty)) => pairself (Thm.cterm_of thy) | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 311 | (Var ((v, 0), map_atyps (fn _ => TVar ((Name.aT, 0), sort)) ty), | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 312 | Const (c, map_atyps (fn _ => TVar ((Name.aT, 0), sort)) ty))) param_map); | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 313 | val axiom_premises = map_filter (#axiom o the_class_data thy) sups | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 314 | @ the_list assm_axiom; | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 315 | in intro | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 316 | |> instantiate thy [class] | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 317 | |> (fn thm => thm OF axiom_premises) | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 318 | |> Drule.standard' | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 319 | |> Thm.close_derivation | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 320 | |> SOME | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 321 | end; | 
| 25024 | 322 | |
| 28715 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 323 | fun calculate_rules thy phi sups base_sort param_map axiom class = | 
| 25062 | 324 | let | 
| 25711 | 325 | fun instantiate thy sort = Thm.instantiate ([pairself (Thm.ctyp_of thy o TVar o pair (Name.aT, 0)) | 
| 326 | (base_sort, sort)], map (fn (v, (c, ty)) => pairself (Thm.cterm_of thy) | |
| 327 | (Var ((v, 0), map_atyps (fn _ => TVar ((Name.aT, 0), sort)) ty), | |
| 328 | Const (c, map_atyps (fn _ => TVar ((Name.aT, 0), sort)) ty))) param_map); | |
| 329 | val defs = these_defs thy sups; | |
| 28715 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 330 | val assm_intro = Locale.intros thy class | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 331 | |> fst | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 332 | |> map (instantiate thy base_sort) | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 333 | |> map (MetaSimplifier.rewrite_rule defs) | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 334 | |> map Thm.close_derivation | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 335 | |> try the_single; | 
| 25711 | 336 | val fixate = Thm.instantiate | 
| 337 | (map (pairself (Thm.ctyp_of thy)) [(TVar ((Name.aT, 0), []), TFree (Name.aT, base_sort)), | |
| 338 | (TVar ((Name.aT, 0), base_sort), TFree (Name.aT, base_sort))], []) | |
| 25618 | 339 | val of_class_sups = if null sups | 
| 25711 | 340 | then map (fixate o Thm.class_triv thy) base_sort | 
| 341 | else map (fixate o #of_class o the_class_data thy) sups; | |
| 25683 | 342 | val locale_dests = map Drule.standard' (Locale.dests thy class); | 
| 25711 | 343 | val num_trivs = case length locale_dests | 
| 344 | of 0 => if is_none axiom then 0 else 1 | |
| 345 | | n => n; | |
| 346 | val pred_trivs = if num_trivs = 0 then [] | |
| 347 | else the axiom | |
| 348 | |> Thm.prop_of | |
| 349 | |> (map_types o map_atyps o K) (TFree (Name.aT, base_sort)) | |
| 350 | |> (Thm.assume o Thm.cterm_of thy) | |
| 351 | |> replicate num_trivs; | |
| 28715 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 352 | val axclass_intro = (#intro o AxClass.get_info thy) class; | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 353 | val of_class = (fixate axclass_intro OF of_class_sups OF locale_dests OF pred_trivs) | 
| 25711 | 354 | |> Drule.standard' | 
| 26628 
63306cb94313
replaced Drule.close_derivation/Goal.close_result by Thm.close_derivation (removed obsolete compression);
 wenzelm parents: 
26596diff
changeset | 355 | |> Thm.close_derivation; | 
| 28715 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 356 | in (assm_intro, of_class) end; | 
| 24218 | 357 | |
| 27684 
f45fd1159a4b
subclass now also works for subclasses with empty specificaton
 haftmann parents: 
27281diff
changeset | 358 | fun prove_subclass (sub, sup) some_thm thy = | 
| 25618 | 359 | let | 
| 25711 | 360 | val of_class = (#of_class o the_class_data thy) sup; | 
| 27684 
f45fd1159a4b
subclass now also works for subclasses with empty specificaton
 haftmann parents: 
27281diff
changeset | 361 | val intro = case some_thm | 
| 
f45fd1159a4b
subclass now also works for subclasses with empty specificaton
 haftmann parents: 
27281diff
changeset | 362 | of SOME thm => Drule.standard' (of_class OF [Drule.standard' thm]) | 
| 
f45fd1159a4b
subclass now also works for subclasses with empty specificaton
 haftmann parents: 
27281diff
changeset | 363 | | NONE => Thm.instantiate ([pairself (Thm.ctyp_of thy o TVar o pair (Name.aT, 0)) | 
| 
f45fd1159a4b
subclass now also works for subclasses with empty specificaton
 haftmann parents: 
27281diff
changeset | 364 | ([], [sub])], []) of_class; | 
| 
f45fd1159a4b
subclass now also works for subclasses with empty specificaton
 haftmann parents: 
27281diff
changeset | 365 | val classrel = (intro OF (the_list o #axiom o the_class_data thy) sub) | 
| 
f45fd1159a4b
subclass now also works for subclasses with empty specificaton
 haftmann parents: 
27281diff
changeset | 366 | |> Thm.close_derivation; | 
| 25618 | 367 | in | 
| 368 | thy | |
| 369 | |> AxClass.add_classrel classrel | |
| 27684 
f45fd1159a4b
subclass now also works for subclasses with empty specificaton
 haftmann parents: 
27281diff
changeset | 370 | |> prove_interpretation_in (ALLGOALS (ProofContext.fact_tac (the_list some_thm))) | 
| 25618 | 371 | I (sub, Locale.Locale sup) | 
| 372 | |> ClassData.map (Graph.add_edge (sub, sup)) | |
| 373 | end; | |
| 374 | ||
| 28715 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 375 | fun note_assm_intro class assm_intro thy = | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 376 | thy | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 377 | |> Sign.sticky_prefix (class_prefix class ^ "_" ^ AxClass.axiomsN) | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 378 | |> PureThy.store_thm (AxClass.introN, assm_intro) | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 379 | |> snd | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 380 | |> Sign.restore_naming thy; | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 381 | |
| 24218 | 382 | fun intro_classes_tac facts st = | 
| 383 | let | |
| 384 | val thy = Thm.theory_of_thm st; | |
| 385 | val classes = Sign.all_classes thy; | |
| 386 | val class_trivs = map (Thm.class_triv thy) classes; | |
| 25618 | 387 | val class_intros = map_filter (try (#intro o AxClass.get_info thy)) classes; | 
| 388 | val assm_intros = these_assm_intros thy; | |
| 24218 | 389 | in | 
| 25618 | 390 | Method.intros_tac (class_trivs @ class_intros @ assm_intros) facts st | 
| 24218 | 391 | end; | 
| 392 | ||
| 26470 | 393 | fun default_intro_tac ctxt [] = | 
| 394 | intro_classes_tac [] ORELSE Locale.intro_locales_tac true ctxt [] | |
| 395 | | default_intro_tac _ _ = no_tac; | |
| 24218 | 396 | |
| 397 | fun default_tac rules ctxt facts = | |
| 398 | HEADGOAL (Method.some_rule_tac rules ctxt facts) ORELSE | |
| 26470 | 399 | default_intro_tac ctxt facts; | 
| 24218 | 400 | |
| 26463 | 401 | val _ = Context.>> (Context.map_theory | 
| 402 | (Method.add_methods | |
| 403 |    [("intro_classes", Method.no_args (Method.METHOD intro_classes_tac),
 | |
| 404 | "back-chain introduction rules of classes"), | |
| 405 |     ("default", Method.thms_ctxt_args (Method.METHOD oo default_tac),
 | |
| 406 | "apply some intro/elim rule")])); | |
| 407 | ||
| 24218 | 408 | |
| 24589 | 409 | (** classes and class target **) | 
| 24218 | 410 | |
| 25002 | 411 | (* class context syntax *) | 
| 24748 | 412 | |
| 26238 | 413 | fun synchronize_class_syntax sups base_sort ctxt = | 
| 24914 
95cda5dd58d5
added proper subclass concept; improved class target
 haftmann parents: 
24901diff
changeset | 414 | let | 
| 25344 
00c2179db769
synchronize_syntax: improved declare_const (still inactive);
 wenzelm parents: 
25326diff
changeset | 415 | val thy = ProofContext.theory_of ctxt; | 
| 26596 | 416 | val algebra = Sign.classes_of thy; | 
| 25083 | 417 | val operations = these_operations thy sups; | 
| 26518 | 418 | fun subst_class_typ sort = map_type_tfree (K (TVar ((Name.aT, 0), sort))); | 
| 419 | val primary_constraints = | |
| 25368 
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
 haftmann parents: 
25344diff
changeset | 420 | (map o apsnd) (subst_class_typ base_sort o fst o snd) operations; | 
| 26518 | 421 | val secondary_constraints = | 
| 25368 
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
 haftmann parents: 
25344diff
changeset | 422 | (map o apsnd) (fn (class, (ty, _)) => subst_class_typ [class] ty) operations; | 
| 25318 
c8352b38d47d
synchronize_syntax: declare operations within the local scope of fixes/consts
 wenzelm parents: 
25311diff
changeset | 423 | fun declare_const (c, _) = | 
| 
c8352b38d47d
synchronize_syntax: declare operations within the local scope of fixes/consts
 wenzelm parents: 
25311diff
changeset | 424 | let val b = Sign.base_name c | 
| 25344 
00c2179db769
synchronize_syntax: improved declare_const (still inactive);
 wenzelm parents: 
25326diff
changeset | 425 | in Sign.intern_const thy b = c ? Variable.declare_const (b, c) end; | 
| 26518 | 426 | fun improve (c, ty) = (case AList.lookup (op =) primary_constraints c | 
| 26238 | 427 | of SOME ty' => (case try (Type.raw_match (ty', ty)) Vartab.empty | 
| 428 | of SOME tyenv => (case Vartab.lookup tyenv (Name.aT, 0) | |
| 26596 | 429 | of SOME (_, ty' as TVar (tvar as (vi, sort))) => | 
| 26238 | 430 | if TypeInfer.is_param vi | 
| 26596 | 431 | andalso Sorts.sort_le algebra (base_sort, sort) | 
| 432 | then SOME (ty', TFree (Name.aT, base_sort)) | |
| 433 | else NONE | |
| 26238 | 434 | | _ => NONE) | 
| 435 | | NONE => NONE) | |
| 436 | | NONE => NONE) | |
| 437 | fun subst (c, ty) = Option.map snd (AList.lookup (op =) operations c); | |
| 25368 
f12613fda79d
proper implementation of check phase; non-qualified names for class operations
 haftmann parents: 
25344diff
changeset | 438 | val unchecks = map (fn (c, (_, (ty, t))) => (t, Const (c, ty))) operations; | 
| 25083 | 439 | in | 
| 440 | ctxt | |
| 26518 | 441 | |> fold declare_const primary_constraints | 
| 442 | |> Overloading.map_improvable_syntax (K (((primary_constraints, secondary_constraints), | |
| 26730 | 443 | (((improve, subst), true), unchecks)), false)) | 
| 26518 | 444 | |> Overloading.set_primary_constraints | 
| 25083 | 445 | end; | 
| 446 | ||
| 447 | fun refresh_syntax class ctxt = | |
| 25002 | 448 | let | 
| 449 | val thy = ProofContext.theory_of ctxt; | |
| 25062 | 450 | val base_sort = (#base_sort o the_class_data thy) class; | 
| 26238 | 451 | in synchronize_class_syntax [class] base_sort ctxt end; | 
| 25002 | 452 | |
| 25344 
00c2179db769
synchronize_syntax: improved declare_const (still inactive);
 wenzelm parents: 
25326diff
changeset | 453 | fun init_ctxt sups base_sort ctxt = | 
| 25083 | 454 | ctxt | 
| 455 | |> Variable.declare_term | |
| 456 | (Logic.mk_type (TFree (Name.aT, base_sort))) | |
| 26238 | 457 | |> synchronize_class_syntax sups base_sort | 
| 458 | |> Overloading.add_improvable_syntax; | |
| 24901 
d3cbf79769b9
added first version of user-space type system for class target
 haftmann parents: 
24847diff
changeset | 459 | |
| 25311 | 460 | fun init class thy = | 
| 461 | thy | |
| 462 | |> Locale.init class | |
| 25344 
00c2179db769
synchronize_syntax: improved declare_const (still inactive);
 wenzelm parents: 
25326diff
changeset | 463 | |> init_ctxt [class] ((#base_sort o the_class_data thy) class); | 
| 24914 
95cda5dd58d5
added proper subclass concept; improved class target
 haftmann parents: 
24901diff
changeset | 464 | |
| 24748 | 465 | |
| 27690 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 466 | (* class target *) | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 467 | |
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 468 | fun declare class pos ((c, mx), dict) thy = | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 469 | let | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 470 | val prfx = class_prefix class; | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 471 | val thy' = thy |> Sign.add_path prfx; | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 472 | val phi = morphism thy' class; | 
| 28715 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 473 | val inst = the_inst thy' class; | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 474 | val params = map (apsnd fst o snd) (these_params thy' [class]); | 
| 27690 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 475 | |
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 476 | val c' = Sign.full_name thy' c; | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 477 | val dict' = Morphism.term phi dict; | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 478 | val dict_def = map_types Logic.unvarifyT dict'; | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 479 | val ty' = Term.fastype_of dict_def; | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 480 | val ty'' = Type.strip_sorts ty'; | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 481 | val def_eq = Logic.mk_equals (Const (c', ty'), dict_def); | 
| 28674 | 482 | fun get_axiom thy = ((Thm.varifyT o Thm.axiom thy) c', thy); | 
| 27690 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 483 | in | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 484 | thy' | 
| 28110 | 485 | |> Sign.declare_const pos ((Name.binding c, ty''), mx) |> snd | 
| 27690 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 486 | |> Thm.add_def false false (c, def_eq) | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 487 | |>> Thm.symmetric | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 488 | ||>> get_axiom | 
| 28715 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 489 | |-> (fn (def, def') => prove_class_interpretation class inst params [] [def] | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 490 | #> snd | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 491 | (*assumption: interpretation cookie does not change | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 492 | by adding equations to interpretation*) | 
| 27690 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 493 | #> register_operation class (c', (dict', SOME (Thm.symmetric def'))) | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 494 | #> PureThy.store_thm (c ^ "_raw", def') | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 495 | #> snd) | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 496 | |> Sign.restore_naming thy | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 497 | |> Sign.add_const_constraint (c', SOME ty') | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 498 | end; | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 499 | |
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 500 | fun abbrev class prmode pos ((c, mx), rhs) thy = | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 501 | let | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 502 | val prfx = class_prefix class; | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 503 | val thy' = thy |> Sign.add_path prfx; | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 504 | |
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 505 | val unchecks = map (fn (c, (_, (ty, t))) => (t, Const (c, ty))) | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 506 | (these_operations thy [class]); | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 507 | val c' = Sign.full_name thy' c; | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 508 | val rhs' = Pattern.rewrite_term thy unchecks [] rhs; | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 509 | val rhs'' = map_types Logic.varifyT rhs'; | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 510 | val ty' = Term.fastype_of rhs'; | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 511 | in | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 512 | thy' | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 513 | |> Sign.add_abbrev (#1 prmode) pos (c, map_types Type.strip_sorts rhs'') |> snd | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 514 | |> Sign.add_const_constraint (c', SOME ty') | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 515 | |> Sign.notation true prmode [(Const (c', ty'), mx)] | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 516 | |> not (#1 prmode = PrintMode.input) ? register_operation class (c', (rhs', NONE)) | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 517 | |> Sign.restore_naming thy | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 518 | end; | 
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 519 | |
| 
24738db98d34
some steps towards explicit class target for canonical interpretation
 haftmann parents: 
27684diff
changeset | 520 | |
| 24589 | 521 | (* class definition *) | 
| 24218 | 522 | |
| 523 | local | |
| 524 | ||
| 26247 | 525 | fun gen_class_spec prep_class process_expr thy raw_supclasses raw_elems = | 
| 24218 | 526 | let | 
| 24748 | 527 | val supclasses = map (prep_class thy) raw_supclasses; | 
| 528 | val supsort = Sign.minimize_sort thy supclasses; | |
| 25618 | 529 | val sups = filter (is_class thy) supsort; | 
| 26995 | 530 | val supparam_names = map fst (these_params thy sups); | 
| 531 | val _ = if has_duplicates (op =) supparam_names | |
| 532 |       then error ("Duplicate parameter(s) in superclasses: "
 | |
| 533 | ^ (commas o map quote o duplicates (op =)) supparam_names) | |
| 534 | else (); | |
| 25618 | 535 | val base_sort = if null sups then supsort else | 
| 26167 | 536 | foldr1 (Sorts.inter_sort (Sign.classes_of thy)) | 
| 537 | (map (#base_sort o the_class_data thy) sups); | |
| 25038 | 538 | val suplocales = map Locale.Locale sups; | 
| 24748 | 539 | val supexpr = Locale.Merge suplocales; | 
| 540 | val supparams = (map fst o Locale.parameters_of_expr thy) supexpr; | |
| 26247 | 541 | val mergeexpr = Locale.Merge (suplocales); | 
| 24748 | 542 | val constrain = Element.Constrains ((map o apsnd o map_atyps) | 
| 26167 | 543 | (K (TFree (Name.aT, base_sort))) supparams); | 
| 25683 | 544 | fun fork_syn (Element.Fixes xs) = | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28017diff
changeset | 545 | fold_map (fn (c, ty, syn) => cons (Name.name_of c, syn) #> pair (c, ty, NoSyn)) xs | 
| 25683 | 546 | #>> Element.Fixes | 
| 547 | | fork_syn x = pair x; | |
| 548 | fun fork_syntax elems = | |
| 549 | let | |
| 550 | val (elems', global_syntax) = fold_map fork_syn elems []; | |
| 26247 | 551 | in (constrain :: elems', global_syntax) end; | 
| 25683 | 552 | val (elems, global_syntax) = | 
| 553 | ProofContext.init thy | |
| 554 | |> Locale.cert_expr supexpr [constrain] | |
| 555 | |> snd | |
| 556 | |> init_ctxt sups base_sort | |
| 557 | |> process_expr Locale.empty raw_elems | |
| 558 | |> fst | |
| 559 | |> fork_syntax | |
| 560 | in (((sups, supparams), (supsort, base_sort, mergeexpr)), (elems, global_syntax)) end; | |
| 24748 | 561 | |
| 26247 | 562 | val read_class_spec = gen_class_spec Sign.intern_class Locale.read_expr; | 
| 563 | val check_class_spec = gen_class_spec (K I) Locale.cert_expr; | |
| 24748 | 564 | |
| 28715 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 565 | fun add_consts bname class base_sort sups supparams global_syntax thy = | 
| 24968 
f9bafc868847
replaced Sign.add_consts_authentic by Sign.declare_const;
 wenzelm parents: 
24949diff
changeset | 566 | let | 
| 25683 | 567 | val supconsts = map fst supparams | 
| 26518 | 568 | |> AList.make (snd o the o AList.lookup (op =) (these_params thy sups)) | 
| 25683 | 569 | |> (map o apsnd o apsnd o map_atyps o K o TFree) (Name.aT, [class]); | 
| 570 | val all_params = map fst (Locale.parameters_of thy class); | |
| 28715 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 571 | val raw_params = (snd o chop (length supparams)) all_params; | 
| 25683 | 572 | fun add_const (v, raw_ty) thy = | 
| 573 | let | |
| 574 | val c = Sign.full_name thy v; | |
| 575 | val ty = map_atyps (K (TFree (Name.aT, base_sort))) raw_ty; | |
| 576 | val ty0 = Type.strip_sorts ty; | |
| 577 | val ty' = map_atyps (K (TFree (Name.aT, [class]))) ty0; | |
| 578 | val syn = (the_default NoSyn o AList.lookup (op =) global_syntax) v; | |
| 579 | in | |
| 580 | thy | |
| 28110 | 581 | |> Sign.declare_const [] ((Name.binding v, ty0), syn) | 
| 25683 | 582 | |> snd | 
| 583 | |> pair ((v, ty), (c, ty')) | |
| 584 | end; | |
| 28715 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 585 | in | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 586 | thy | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 587 | |> Sign.add_path (Logic.const_of_class bname) | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 588 | |> fold_map add_const raw_params | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 589 | ||> Sign.restore_naming thy | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 590 | |-> (fn params => pair (supconsts @ (map o apfst) fst params, params)) | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 591 | end; | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 592 | |
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 593 | fun adjungate_axclass bname class base_sort sups supsort supparams global_syntax thy = | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 594 | let | 
| 25683 | 595 | fun globalize param_map = map_aterms | 
| 596 | (fn Free (v, ty) => Const ((fst o the o AList.lookup (op =) param_map) v, ty) | |
| 597 | | t => t); | |
| 598 | val raw_pred = Locale.intros thy class | |
| 599 | |> fst | |
| 600 | |> map (Logic.unvarify o Logic.strip_imp_concl o Thm.prop_of); | |
| 601 | fun get_axiom thy = case (#axioms o AxClass.get_info thy) class | |
| 602 | of [] => NONE | |
| 603 | | [thm] => SOME thm; | |
| 24968 
f9bafc868847
replaced Sign.add_consts_authentic by Sign.declare_const;
 wenzelm parents: 
24949diff
changeset | 604 | in | 
| 
f9bafc868847
replaced Sign.add_consts_authentic by Sign.declare_const;
 wenzelm parents: 
24949diff
changeset | 605 | thy | 
| 28715 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 606 | |> add_consts bname class base_sort sups supparams global_syntax | 
| 25683 | 607 | |-> (fn (param_map, params) => AxClass.define_class (bname, supsort) | 
| 26518 | 608 | (map (fst o snd) params) | 
| 28083 
103d9282a946
explicit type Name.binding for higher-specification elements;
 wenzelm parents: 
28017diff
changeset | 609 | [((Name.binding (bname ^ "_" ^ AxClass.axiomsN), []), map (globalize param_map) raw_pred)] | 
| 25683 | 610 | #> snd | 
| 611 | #> `get_axiom | |
| 612 | #-> (fn assm_axiom => fold (Sign.add_const_constraint o apsnd SOME o snd) params | |
| 28715 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 613 | #> pair (map (Const o snd) param_map, param_map, params, assm_axiom))) | 
| 24968 
f9bafc868847
replaced Sign.add_consts_authentic by Sign.declare_const;
 wenzelm parents: 
24949diff
changeset | 614 | end; | 
| 
f9bafc868847
replaced Sign.add_consts_authentic by Sign.declare_const;
 wenzelm parents: 
24949diff
changeset | 615 | |
| 26518 | 616 | fun gen_class prep_spec bname raw_supclasses raw_elems thy = | 
| 24748 | 617 | let | 
| 25038 | 618 | val class = Sign.full_name thy bname; | 
| 25683 | 619 | val (((sups, supparams), (supsort, base_sort, mergeexpr)), (elems, global_syntax)) = | 
| 26247 | 620 | prep_spec thy raw_supclasses raw_elems; | 
| 28715 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 621 | val supconsts = map (apsnd fst o snd) (these_params thy sups); | 
| 24218 | 622 | in | 
| 623 | thy | |
| 27684 
f45fd1159a4b
subclass now also works for subclasses with empty specificaton
 haftmann parents: 
27281diff
changeset | 624 | |> Locale.add_locale_i "" bname mergeexpr elems | 
| 25038 | 625 | |> snd | 
| 25311 | 626 | |> ProofContext.theory_of | 
| 26518 | 627 | |> adjungate_axclass bname class base_sort sups supsort supparams global_syntax | 
| 28715 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 628 | |-> (fn (inst, param_map, params, assm_axiom) => | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 629 | `(fn thy => calculate_axiom thy sups base_sort assm_axiom param_map class) | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 630 | #-> (fn axiom => | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 631 | prove_class_interpretation class inst | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 632 | (supconsts @ map (pair class o fst o snd) params) (the_list axiom) [] | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 633 | #-> (fn morphism => | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 634 | `(fn thy => activate_class_morphism thy class inst morphism) | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 635 | #-> (fn phi => | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 636 | `(fn thy => calculate_rules thy phi sups base_sort param_map axiom class) | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 637 | #-> (fn (assm_intro, of_class) => | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 638 | add_class_data ((class, sups), (params, base_sort, inst, | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 639 | morphism, axiom, assm_intro, of_class)) | 
| 
238f9966c80e
class morphism stemming from locale interpretation
 haftmann parents: 
28674diff
changeset | 640 | #> fold (note_assm_intro class) (the_list assm_intro)))))) | 
| 25268 | 641 | |> init class | 
| 25038 | 642 | |> pair class | 
| 24218 | 643 | end; | 
| 644 | ||
| 645 | in | |
| 646 | ||
| 26518 | 647 | val class_cmd = gen_class read_class_spec; | 
| 648 | val class = gen_class check_class_spec; | |
| 24218 | 649 | |
| 650 | end; (*local*) | |
| 651 | ||
| 652 | ||
| 25462 | 653 | |
| 654 | (** instantiation target **) | |
| 655 | ||
| 656 | (* bookkeeping *) | |
| 657 | ||
| 658 | datatype instantiation = Instantiation of {
 | |
| 25864 | 659 | arities: string list * (string * sort) list * sort, | 
| 25462 | 660 | params: ((string * string) * (string * typ)) list | 
| 25603 | 661 | (*(instantiation parameter, type constructor), (local instantiation parameter, typ)*) | 
| 25462 | 662 | } | 
| 663 | ||
| 664 | structure Instantiation = ProofDataFun | |
| 665 | ( | |
| 666 | type T = instantiation | |
| 25536 | 667 |   fun init _ = Instantiation { arities = ([], [], []), params = [] };
 | 
| 25462 | 668 | ); | 
| 669 | ||
| 25485 | 670 | fun mk_instantiation (arities, params) = | 
| 671 |   Instantiation { arities = arities, params = params };
 | |
| 25514 | 672 | fun get_instantiation lthy = case Instantiation.get (LocalTheory.target_of lthy) | 
| 25485 | 673 | of Instantiation data => data; | 
| 25514 | 674 | fun map_instantiation f = (LocalTheory.target o Instantiation.map) | 
| 675 |   (fn Instantiation { arities, params } => mk_instantiation (f (arities, params)));
 | |
| 25462 | 676 | |
| 25514 | 677 | fun the_instantiation lthy = case get_instantiation lthy | 
| 25536 | 678 |  of { arities = ([], [], []), ... } => error "No instantiation target"
 | 
| 25485 | 679 | | data => data; | 
| 25462 | 680 | |
| 25485 | 681 | val instantiation_params = #params o get_instantiation; | 
| 25462 | 682 | |
| 25514 | 683 | fun instantiation_param lthy v = instantiation_params lthy | 
| 25462 | 684 | |> find_first (fn (_, (v', _)) => v = v') | 
| 685 | |> Option.map (fst o fst); | |
| 686 | ||
| 687 | ||
| 688 | (* syntax *) | |
| 689 | ||
| 26238 | 690 | fun synchronize_inst_syntax ctxt = | 
| 25462 | 691 | let | 
| 26259 | 692 |     val Instantiation { arities = (_, _, sort), params = params } = Instantiation.get ctxt;
 | 
| 26238 | 693 | val thy = ProofContext.theory_of ctxt; | 
| 694 | fun subst (c, ty) = case AxClass.inst_tyco_of thy (c, ty) | |
| 695 | of SOME tyco => (case AList.lookup (op =) params (c, tyco) | |
| 696 | of SOME (v_ty as (_, ty)) => SOME (ty, Free v_ty) | |
| 697 | | NONE => NONE) | |
| 698 | | NONE => NONE; | |
| 699 | val unchecks = | |
| 700 | map (fn ((c, _), v_ty as (_, ty)) => (Free v_ty, Const (c, ty))) params; | |
| 701 | in | |
| 702 | ctxt | |
| 26259 | 703 | |> Overloading.map_improvable_syntax | 
| 26730 | 704 | (fn (((primary_constraints, _), (((improve, _), _), _)), _) => | 
| 705 | (((primary_constraints, []), (((improve, subst), false), unchecks)), false)) | |
| 26238 | 706 | end; | 
| 25462 | 707 | |
| 708 | ||
| 709 | (* target *) | |
| 710 | ||
| 25485 | 711 | val sanatize_name = (*FIXME*) | 
| 712 | let | |
| 25574 | 713 | fun is_valid s = Symbol.is_ascii_letter s orelse Symbol.is_ascii_digit s | 
| 714 | orelse s = "'" orelse s = "_"; | |
| 25485 | 715 | val is_junk = not o is_valid andf Symbol.is_regular; | 
| 716 | val junk = Scan.many is_junk; | |
| 717 | val scan_valids = Symbol.scanner "Malformed input" | |
| 718 | ((junk |-- | |
| 719 | (Scan.optional (Scan.one Symbol.is_ascii_letter) "x" ^^ (Scan.many is_valid >> implode) | |
| 720 | --| junk)) | |
| 25999 | 721 | ::: Scan.repeat ((Scan.many1 is_valid >> implode) --| junk)); | 
| 25485 | 722 | in | 
| 723 | explode #> scan_valids #> implode | |
| 724 | end; | |
| 725 | ||
| 26259 | 726 | fun type_name "*" = "prod" | 
| 727 | | type_name "+" = "sum" | |
| 728 | | type_name s = sanatize_name (NameSpace.base s); (*FIXME*) | |
| 729 | ||
| 26518 | 730 | fun resort_terms pp algebra consts constraints ts = | 
| 731 | let | |
| 732 | fun matchings (Const (c_ty as (c, _))) = (case constraints c | |
| 733 | of NONE => I | |
| 734 | | SOME sorts => fold2 (curry (Sorts.meet_sort algebra)) | |
| 735 | (Consts.typargs consts c_ty) sorts) | |
| 736 | | matchings _ = I | |
| 737 | val tvartab = (fold o fold_aterms) matchings ts Vartab.empty | |
| 26642 
454d11701fa4
Sorts.class_error: produce message only (formerly msg_class_error);
 wenzelm parents: 
26628diff
changeset | 738 | handle Sorts.CLASS_ERROR e => error (Sorts.class_error pp e); | 
| 27089 
480f19078b65
fixed wrong treatment of type variables in instantiation target
 haftmann parents: 
26995diff
changeset | 739 | val inst = map_type_tvar | 
| 
480f19078b65
fixed wrong treatment of type variables in instantiation target
 haftmann parents: 
26995diff
changeset | 740 | (fn (vi, sort) => TVar (vi, the_default sort (Vartab.lookup tvartab vi))); | 
| 26518 | 741 | in if Vartab.is_empty tvartab then NONE else SOME ((map o map_types) inst ts) end; | 
| 742 | ||
| 25864 | 743 | fun init_instantiation (tycos, vs, sort) thy = | 
| 25462 | 744 | let | 
| 25536 | 745 | val _ = if null tycos then error "At least one arity must be given" else (); | 
| 26518 | 746 | val params = these_params thy sort; | 
| 747 | fun get_param tyco (param, (_, (c, ty))) = if can (AxClass.param_of_inst thy) (c, tyco) | |
| 25603 | 748 | then NONE else SOME ((c, tyco), | 
| 25864 | 749 | (param ^ "_" ^ type_name tyco, map_atyps (K (Type (tyco, map TFree vs))) ty)); | 
| 26518 | 750 | val inst_params = map_product get_param tycos params |> map_filter I; | 
| 751 | val primary_constraints = map (apsnd | |
| 752 | (map_atyps (K (TVar ((Name.aT, 0), [])))) o snd o snd) params; | |
| 26939 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 wenzelm parents: 
26761diff
changeset | 753 | val pp = Syntax.pp_global thy; | 
| 26518 | 754 | val algebra = Sign.classes_of thy | 
| 755 | |> fold (fn tyco => Sorts.add_arities pp | |
| 756 | (tyco, map (fn class => (class, map snd vs)) sort)) tycos; | |
| 757 | val consts = Sign.consts_of thy; | |
| 758 | val improve_constraints = AList.lookup (op =) | |
| 759 | (map (fn (_, (class, (c, _))) => (c, [[class]])) params); | |
| 760 | fun resort_check ts ctxt = case resort_terms pp algebra consts improve_constraints ts | |
| 761 | of NONE => NONE | |
| 762 | | SOME ts' => SOME (ts', ctxt); | |
| 763 | fun improve (c, ty) = case AxClass.inst_tyco_of thy (c, ty) | |
| 26329 
3e58e4c67a2a
instantiation less liberal with dangling constraints
 haftmann parents: 
26259diff
changeset | 764 | of SOME tyco => (case AList.lookup (op =) inst_params (c, tyco) | 
| 26518 | 765 | of SOME (_, ty') => if Type.raw_instance (ty', ty) then SOME (ty, ty') else NONE | 
| 26259 | 766 | | NONE => NONE) | 
| 767 | | NONE => NONE; | |
| 25485 | 768 | in | 
| 769 | thy | |
| 770 | |> ProofContext.init | |
| 26329 
3e58e4c67a2a
instantiation less liberal with dangling constraints
 haftmann parents: 
26259diff
changeset | 771 | |> Instantiation.put (mk_instantiation ((tycos, vs, sort), inst_params)) | 
| 27281 | 772 | |> fold (Variable.declare_typ o TFree) vs | 
| 26329 
3e58e4c67a2a
instantiation less liberal with dangling constraints
 haftmann parents: 
26259diff
changeset | 773 | |> fold (Variable.declare_names o Free o snd) inst_params | 
| 26259 | 774 | |> (Overloading.map_improvable_syntax o apfst) | 
| 26329 
3e58e4c67a2a
instantiation less liberal with dangling constraints
 haftmann parents: 
26259diff
changeset | 775 | (fn ((_, _), ((_, subst), unchecks)) => | 
| 26730 | 776 | ((primary_constraints, []), (((improve, K NONE), false), []))) | 
| 26259 | 777 | |> Overloading.add_improvable_syntax | 
| 26518 | 778 | |> Context.proof_map (Syntax.add_term_check 0 "resorting" resort_check) | 
| 26238 | 779 | |> synchronize_inst_syntax | 
| 25485 | 780 | end; | 
| 781 | ||
| 26238 | 782 | fun confirm_declaration c = (map_instantiation o apsnd) | 
| 783 | (filter_out (fn (_, (c', _)) => c' = c)) | |
| 784 | #> LocalTheory.target synchronize_inst_syntax | |
| 785 | ||
| 25485 | 786 | fun gen_instantiation_instance do_proof after_qed lthy = | 
| 787 | let | |
| 25864 | 788 | val (tycos, vs, sort) = (#arities o the_instantiation) lthy; | 
| 789 | val arities_proof = maps (fn tyco => Logic.mk_arities (tyco, map snd vs, sort)) tycos; | |
| 25462 | 790 | fun after_qed' results = | 
| 791 | LocalTheory.theory (fold (AxClass.add_arity o Thm.varifyT) results) | |
| 792 | #> after_qed; | |
| 793 | in | |
| 794 | lthy | |
| 795 | |> do_proof after_qed' arities_proof | |
| 796 | end; | |
| 797 | ||
| 25485 | 798 | val instantiation_instance = gen_instantiation_instance (fn after_qed => fn ts => | 
| 25462 | 799 | Proof.theorem_i NONE (after_qed o map the_single) (map (fn t => [(t, [])]) ts)); | 
| 800 | ||
| 25485 | 801 | fun prove_instantiation_instance tac = gen_instantiation_instance (fn after_qed => | 
| 25502 | 802 | fn ts => fn lthy => after_qed (map (fn t => Goal.prove lthy [] [] t | 
| 803 |     (fn {context, ...} => tac context)) ts) lthy) I;
 | |
| 25462 | 804 | |
| 28666 | 805 | fun prove_instantiation_exit tac = prove_instantiation_instance tac | 
| 806 | #> LocalTheory.exit_global; | |
| 807 | ||
| 808 | fun prove_instantiation_exit_result f tac x lthy = | |
| 809 | let | |
| 810 | val phi = ProofContext.export_morphism lthy | |
| 811 | (ProofContext.init (ProofContext.theory_of lthy)); | |
| 812 | val y = f phi x; | |
| 813 | in | |
| 814 | lthy | |
| 815 | |> prove_instantiation_exit (fn ctxt => tac ctxt y) | |
| 816 | |> pair y | |
| 817 | end; | |
| 818 | ||
| 25462 | 819 | fun conclude_instantiation lthy = | 
| 820 | let | |
| 25485 | 821 |     val { arities, params } = the_instantiation lthy;
 | 
| 25864 | 822 | val (tycos, vs, sort) = arities; | 
| 25462 | 823 | val thy = ProofContext.theory_of lthy; | 
| 25597 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 haftmann parents: 
25574diff
changeset | 824 | val _ = map (fn tyco => if Sign.of_sort thy | 
| 25864 | 825 | (Type (tyco, map TFree vs), sort) | 
| 25462 | 826 |       then () else error ("Missing instance proof for type " ^ quote (Sign.extern_type thy tyco)))
 | 
| 25597 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 haftmann parents: 
25574diff
changeset | 827 | tycos; | 
| 
34860182b250
moved instance parameter management from class.ML to axclass.ML
 haftmann parents: 
25574diff
changeset | 828 | in lthy end; | 
| 25462 | 829 | |
| 25603 | 830 | fun pretty_instantiation lthy = | 
| 831 | let | |
| 832 |     val { arities, params } = the_instantiation lthy;
 | |
| 25864 | 833 | val (tycos, vs, sort) = arities; | 
| 25603 | 834 | val thy = ProofContext.theory_of lthy; | 
| 25864 | 835 | fun pr_arity tyco = Syntax.pretty_arity lthy (tyco, map snd vs, sort); | 
| 25603 | 836 | fun pr_param ((c, _), (v, ty)) = | 
| 25864 | 837 | (Pretty.block o Pretty.breaks) [Pretty.str v, Pretty.str "==", | 
| 26939 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 wenzelm parents: 
26761diff
changeset | 838 | (Pretty.str o Sign.extern_const thy) c, Pretty.str "::", Syntax.pretty_typ_global thy ty]; | 
| 25603 | 839 | in | 
| 840 | (Pretty.block o Pretty.fbreaks) | |
| 841 | (Pretty.str "instantiation" :: map pr_arity tycos @ map pr_param params) | |
| 842 | end; | |
| 843 | ||
| 24218 | 844 | end; | 
| 25683 | 845 |