| author | haftmann | 
| Wed, 15 Nov 2006 17:05:38 +0100 | |
| changeset 21379 | a0561695167a | 
| parent 21359 | 072e83a0b5bb | 
| child 21444 | 8f71e2b3fd92 | 
| permissions | -rw-r--r-- | 
| 18168 | 1 | (* Title: Pure/Tools/class_package.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Florian Haftmann, TU Muenchen | |
| 4 | ||
| 19468 | 5 | Type classes derived from primitive axclasses and locales. | 
| 18168 | 6 | *) | 
| 7 | ||
| 8 | signature CLASS_PACKAGE = | |
| 9 | sig | |
| 20964 | 10 | val class: bstring -> class list -> Element.context Locale.element list -> theory -> | 
| 11 | string * Proof.context | |
| 12 | val class_i: bstring -> class list -> Element.context_i Locale.element list -> theory -> | |
| 13 | string * Proof.context | |
| 20601 | 14 | val instance_arity: ((bstring * string list) * string) list | 
| 19136 | 15 | -> bstring * Attrib.src list -> ((bstring * Attrib.src list) * string) list | 
| 18575 | 16 | -> theory -> Proof.state | 
| 20601 | 17 | val instance_arity_i: ((bstring * sort list) * sort) list | 
| 19136 | 18 | -> bstring * attribute list -> ((bstring * attribute list) * term) list | 
| 18575 | 19 | -> theory -> Proof.state | 
| 20182 
79c9ff40d760
tactic for prove_instance_arity now gets definition theorems as arguments
 haftmann parents: 
20175diff
changeset | 20 | val prove_instance_arity: (thm list -> tactic) -> ((string * sort list) * sort) list | 
| 19136 | 21 | -> bstring * attribute list -> ((bstring * attribute list) * term) list | 
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 22 | -> theory -> theory | 
| 19150 | 23 | val instance_sort: string * string -> theory -> Proof.state | 
| 24 | val instance_sort_i: class * sort -> theory -> Proof.state | |
| 25 | val prove_instance_sort: tactic -> class * sort -> theory -> theory | |
| 18168 | 26 | |
| 18755 | 27 | val certify_class: theory -> class -> class | 
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 28 | val certify_sort: theory -> sort -> sort | 
| 20385 | 29 | val read_class: theory -> xstring -> class | 
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 30 | val read_sort: theory -> string -> sort | 
| 20465 | 31 | val operational_algebra: theory -> (sort -> sort) * Sorts.algebra | 
| 18702 | 32 | val the_consts_sign: theory -> class -> string * (string * typ) list | 
| 33 | val the_inst_sign: theory -> class * string -> (string * sort) list * (string * typ) list | |
| 20175 | 34 | val assume_arities_of_sort: theory -> ((string * sort list) * sort) list -> typ * sort -> bool | 
| 35 | val assume_arities_thy: theory -> ((string * sort list) * sort) list -> (theory -> 'a) -> 'a | |
| 36 | (*'a must not keep any reference to theory*) | |
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 37 | |
| 18702 | 38 | val print_classes: theory -> unit | 
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 39 | val intro_classes_tac: thm list -> tactic | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 40 | val default_intro_classes_tac: thm list -> tactic | 
| 18168 | 41 | end; | 
| 42 | ||
| 20182 
79c9ff40d760
tactic for prove_instance_arity now gets definition theorems as arguments
 haftmann parents: 
20175diff
changeset | 43 | structure ClassPackage : CLASS_PACKAGE = | 
| 18168 | 44 | struct | 
| 45 | ||
| 46 | ||
| 20455 | 47 | (** theory data **) | 
| 18168 | 48 | |
| 19456 | 49 | datatype class_data = ClassData of {
 | 
| 18575 | 50 | name_locale: string, | 
| 51 | name_axclass: string, | |
| 52 | var: string, | |
| 20843 | 53 | consts: (string * (string * typ)) list | 
| 54 | (*locale parameter ~> toplevel theory constant*), | |
| 55 | operational: bool (* == at least one class operation, | |
| 56 | or at least two operational superclasses *), | |
| 19966 | 57 | propnames: string list | 
| 58 | } * thm list Symtab.table; | |
| 18168 | 59 | |
| 19456 | 60 | fun rep_classdata (ClassData c) = c; | 
| 61 | ||
| 18575 | 62 | structure ClassData = TheoryDataFun ( | 
| 18168 | 63 | struct | 
| 64 | val name = "Pure/classes"; | |
| 19966 | 65 | type T = class_data Symtab.table; | 
| 66 | val empty = Symtab.empty; | |
| 18168 | 67 | val copy = I; | 
| 68 | val extend = I; | |
| 19966 | 69 | fun merge _ = Symtab.join (fn _ => fn (ClassData (classd, instd1), ClassData (_, instd2)) => | 
| 70 | (ClassData (classd, Symtab.merge (K true) (instd1, instd2)))); | |
| 71 | fun print thy data = | |
| 18575 | 72 | let | 
| 19966 | 73 |         fun pretty_class (name, ClassData ({name_locale, name_axclass, var, consts, ...}, _)) =
 | 
| 18575 | 74 | (Pretty.block o Pretty.fbreaks) [ | 
| 75 |             Pretty.str ("class " ^ name ^ ":"),
 | |
| 76 |             Pretty.str ("locale: " ^ name_locale),
 | |
| 77 |             Pretty.str ("axclass: " ^ name_axclass),
 | |
| 78 |             Pretty.str ("class variable: " ^ var),
 | |
| 79 | (Pretty.block o Pretty.fbreaks) ( | |
| 80 | Pretty.str "constants: " | |
| 19280 
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
 haftmann parents: 
19253diff
changeset | 81 | :: map (fn (_, (c, ty)) => Pretty.str (c ^ " :: " ^ Sign.string_of_typ thy ty)) consts | 
| 18575 | 82 | ) | 
| 83 | ] | |
| 84 | in | |
| 19966 | 85 | (Pretty.writeln o Pretty.chunks o map pretty_class o Symtab.dest) data | 
| 18575 | 86 | end; | 
| 18168 | 87 | end | 
| 88 | ); | |
| 89 | ||
| 18708 | 90 | val _ = Context.add_setup ClassData.init; | 
| 18575 | 91 | val print_classes = ClassData.print; | 
| 92 | ||
| 19038 | 93 | |
| 94 | (* queries *) | |
| 95 | ||
| 19966 | 96 | val lookup_class_data = Option.map rep_classdata oo Symtab.lookup o ClassData.get; | 
| 18168 | 97 | |
| 19038 | 98 | fun the_class_data thy class = | 
| 18168 | 99 | case lookup_class_data thy class | 
| 20455 | 100 |     of NONE => error ("undeclared constructive class " ^ quote class)
 | 
| 18168 | 101 | | SOME data => data; | 
| 102 | ||
| 20455 | 103 | fun the_ancestry thy classes = | 
| 18360 | 104 | let | 
| 20455 | 105 | fun proj_class class = | 
| 20843 | 106 | if is_some (lookup_class_data thy class) | 
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 107 | then [class] | 
| 20455 | 108 | else (Sign.certify_sort thy o maps proj_class o Sign.super_classes thy) class; | 
| 19038 | 109 | fun ancestry class anc = | 
| 110 | anc | |
| 20385 | 111 | |> insert (op =) class | 
| 20455 | 112 | |> fold ancestry ((maps proj_class o Sign.super_classes thy) class); | 
| 19038 | 113 | in fold ancestry classes [] end; | 
| 114 | ||
| 20455 | 115 | val the_parm_map = #consts o fst oo the_class_data; | 
| 116 | ||
| 117 | val the_propnames = #propnames o fst oo the_class_data; | |
| 118 | ||
| 19038 | 119 | fun subst_clsvar v ty_subst = | 
| 120 | map_type_tfree (fn u as (w, _) => | |
| 121 | if w = v then ty_subst else TFree u); | |
| 122 | ||
| 18168 | 123 | |
| 19038 | 124 | (* updaters *) | 
| 125 | ||
| 20843 | 126 | fun add_class_data (class, (name_locale, name_axclass, var, consts, operational, propnames)) = | 
| 19957 | 127 | ClassData.map ( | 
| 19966 | 128 |     Symtab.update_new (class, ClassData ({
 | 
| 19957 | 129 | name_locale = name_locale, | 
| 130 | name_axclass = name_axclass, | |
| 131 | var = var, | |
| 19966 | 132 | consts = consts, | 
| 20843 | 133 | operational = operational, | 
| 19966 | 134 | propnames = propnames}, Symtab.empty)) | 
| 19957 | 135 | ); | 
| 19038 | 136 | |
| 19966 | 137 | fun add_inst_def ((class, tyco), thm) = | 
| 138 | ClassData.map ( | |
| 139 | Symtab.map_entry class (fn ClassData (classd, instd) => | |
| 140 | ClassData (classd, Symtab.insert_list eq_thm (tyco, thm) instd)) | |
| 141 | ); | |
| 19038 | 142 | |
| 20106 | 143 | |
| 20385 | 144 | (* certification and reading *) | 
| 19038 | 145 | |
| 146 | fun certify_class thy class = | |
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 147 | (fn class => (the_class_data thy class; class)) (Sign.certify_class thy class); | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 148 | |
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 149 | fun certify_sort thy sort = | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 150 | map (fn class => (the_class_data thy class; class)) (Sign.certify_sort thy sort); | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 151 | |
| 20385 | 152 | fun read_class thy = | 
| 19957 | 153 | certify_class thy o Sign.intern_class thy; | 
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 154 | |
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 155 | fun read_sort thy = | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 156 | certify_sort thy o Sign.read_sort thy; | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 157 | |
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 158 | |
| 20455 | 159 | |
| 160 | (** contexts with arity assumptions **) | |
| 20175 | 161 | |
| 162 | fun assume_arities_of_sort thy arities ty_sort = | |
| 163 | let | |
| 164 | val pp = Sign.pp thy; | |
| 165 | val algebra = Sign.classes_of thy | |
| 166 | |> fold (fn ((tyco, asorts), sort) => | |
| 167 | Sorts.add_arities pp (tyco, map (fn class => (class, asorts)) sort)) arities; | |
| 168 | in Sorts.of_sort algebra ty_sort end; | |
| 169 | ||
| 170 | fun assume_arities_thy thy arities f = | |
| 171 | let | |
| 172 | val thy_read = (fold (fn ((tyco, asorts), sort) | |
| 173 | => Sign.primitive_arity (tyco, asorts, sort)) arities o Theory.copy) thy | |
| 174 | in f thy_read end; | |
| 175 | ||
| 20455 | 176 | |
| 177 | ||
| 178 | (** tactics and methods **) | |
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 179 | |
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 180 | fun intro_classes_tac facts st = | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 181 | (ALLGOALS (Method.insert_tac facts THEN' | 
| 19928 | 182 | REPEAT_ALL_NEW (resolve_tac (AxClass.class_intros (Thm.theory_of_thm st)))) | 
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 183 | THEN Tactic.distinct_subgoals_tac) st; | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 184 | |
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 185 | fun default_intro_classes_tac [] = intro_classes_tac [] | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 186 | | default_intro_classes_tac _ = Tactical.no_tac; (*no error message!*) | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 187 | |
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 188 | fun default_tac rules ctxt facts = | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 189 | HEADGOAL (Method.some_rule_tac rules ctxt facts) ORELSE | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 190 | default_intro_classes_tac facts; | 
| 19038 | 191 | |
| 19468 | 192 | val _ = Context.add_setup (Method.add_methods | 
| 193 |  [("intro_classes", Method.no_args (Method.METHOD intro_classes_tac),
 | |
| 194 | "back-chain introduction rules of classes"), | |
| 195 |   ("default", Method.thms_ctxt_args (Method.METHOD oo default_tac),
 | |
| 196 | "apply some intro/elim rule")]); | |
| 197 | ||
| 19038 | 198 | |
| 20455 | 199 | |
| 200 | (** axclass instances **) | |
| 19246 | 201 | |
| 202 | local | |
| 203 | ||
| 20175 | 204 | fun gen_instance mk_prop add_thm after_qed insts thy = | 
| 20368 
2461a0485623
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 wenzelm parents: 
20353diff
changeset | 205 | let | 
| 
2461a0485623
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 wenzelm parents: 
20353diff
changeset | 206 | fun after_qed' results = | 
| 
2461a0485623
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 wenzelm parents: 
20353diff
changeset | 207 | ProofContext.theory ((fold o fold) add_thm results #> after_qed); | 
| 
2461a0485623
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 wenzelm parents: 
20353diff
changeset | 208 | in | 
| 
2461a0485623
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 wenzelm parents: 
20353diff
changeset | 209 | thy | 
| 
2461a0485623
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 wenzelm parents: 
20353diff
changeset | 210 | |> ProofContext.init | 
| 21359 | 211 | |> Proof.theorem_i PureThy.internalK NONE after_qed' | 
| 212 | ((map (fn t => [(t, [])]) o maps (mk_prop thy)) insts) | |
| 20368 
2461a0485623
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 wenzelm parents: 
20353diff
changeset | 213 | end; | 
| 19246 | 214 | |
| 215 | in | |
| 216 | ||
| 217 | val axclass_instance_arity = | |
| 19412 | 218 | gen_instance (Logic.mk_arities oo Sign.read_arity) AxClass.add_arity; | 
| 19246 | 219 | val axclass_instance_arity_i = | 
| 19412 | 220 | gen_instance (Logic.mk_arities oo Sign.cert_arity) AxClass.add_arity; | 
| 20182 
79c9ff40d760
tactic for prove_instance_arity now gets definition theorems as arguments
 haftmann parents: 
20175diff
changeset | 221 | val axclass_instance_sort = | 
| 
79c9ff40d760
tactic for prove_instance_arity now gets definition theorems as arguments
 haftmann parents: 
20175diff
changeset | 222 | gen_instance (single oo (Logic.mk_classrel oo AxClass.read_classrel)) AxClass.add_classrel I o single; | 
| 19246 | 223 | |
| 224 | end; | |
| 225 | ||
| 226 | ||
| 20455 | 227 | |
| 228 | (** classes and instances **) | |
| 19038 | 229 | |
| 230 | local | |
| 231 | ||
| 19957 | 232 | fun add_axclass_i (name, supsort) params axs thy = | 
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 233 | let | 
| 19395 | 234 | val (c, thy') = thy | 
| 19957 | 235 | |> AxClass.define_class_i (name, supsort) params axs; | 
| 19518 | 236 |     val {intro, axioms, ...} = AxClass.get_definition thy' c;
 | 
| 19395 | 237 | in ((c, (intro, axioms)), thy') end; | 
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 238 | |
| 20385 | 239 | (*FIXME proper locale interface*) | 
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 240 | fun prove_interpretation_i (prfx, atts) expr insts tac thy = | 
| 19038 | 241 | let | 
| 20368 
2461a0485623
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 wenzelm parents: 
20353diff
changeset | 242 | fun ad_hoc_term (Const (c, ty)) = | 
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 243 | let | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 244 | val p = setmp show_types true (setmp show_sorts true (setmp print_mode [] (Sign.pretty_typ thy))) ty; | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 245 | val s = c ^ "::" ^ Pretty.output p; | 
| 20368 
2461a0485623
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 wenzelm parents: 
20353diff
changeset | 246 | in s end | 
| 
2461a0485623
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 wenzelm parents: 
20353diff
changeset | 247 | | ad_hoc_term t = | 
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 248 | let | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 249 | val p = setmp show_types true (setmp show_sorts true (setmp print_mode [] (Sign.pretty_term thy))) t; | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 250 | val s = Pretty.output p; | 
| 20368 
2461a0485623
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 wenzelm parents: 
20353diff
changeset | 251 | in s end; | 
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 252 | in | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 253 | thy | 
| 20368 
2461a0485623
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 wenzelm parents: 
20353diff
changeset | 254 | |> Locale.interpretation I (prfx, atts) expr (map (Option.map ad_hoc_term) insts) | 
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 255 | |> Proof.global_terminal_proof (Method.Basic (fn _ => Method.SIMPLE_METHOD tac), NONE) | 
| 20368 
2461a0485623
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 wenzelm parents: 
20353diff
changeset | 256 | |> ProofContext.theory_of | 
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 257 | end; | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 258 | |
| 19150 | 259 | fun gen_class add_locale prep_class bname raw_supclasses raw_elems thy = | 
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 260 | let | 
| 20428 | 261 | val (elems, exprs) = fold_rev (fn Locale.Elem e => apfst (cons e) | 
| 262 | | Locale.Expr e => apsnd (cons e)) | |
| 263 | raw_elems ([], []); | |
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 264 | val supclasses = map (prep_class thy) raw_supclasses; | 
| 19038 | 265 | val supsort = | 
| 266 | supclasses | |
| 19966 | 267 | |> map (#name_axclass o fst o the_class_data thy) | 
| 19468 | 268 | |> Sign.certify_sort thy | 
| 19038 | 269 | |> null ? K (Sign.defaultS thy); | 
| 20428 | 270 | val expr_supclasses = Locale.Merge | 
| 271 | (map (Locale.Locale o #name_locale o fst o the_class_data thy) supclasses); | |
| 272 | val expr = Locale.Merge (map (Locale.Locale o #name_locale o fst o the_class_data thy) supclasses | |
| 273 | @ exprs); | |
| 19280 
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
 haftmann parents: 
19253diff
changeset | 274 | val mapp_sup = AList.make | 
| 19468 | 275 | (the o AList.lookup (op =) ((flat o map (the_parm_map thy) o the_ancestry thy) supclasses)) | 
| 20428 | 276 | ((map (fst o fst) o Locale.parameters_of_expr thy) expr_supclasses); | 
| 19038 | 277 | fun extract_tyvar_consts thy name_locale = | 
| 278 | let | |
| 20175 | 279 | fun extract_classvar ((c, ty), _) w = | 
| 20428 | 280 | (case Term.add_tfreesT ty [] | 
| 20175 | 281 | of [(v, _)] => (case w | 
| 20385 | 282 |                of SOME u => if u = v then w else error ("Additonal type variable in type signature of constant " ^ quote c)
 | 
| 20175 | 283 | | NONE => SOME v) | 
| 20385 | 284 |             | [] => error ("No type variable in type signature of constant " ^ quote c)
 | 
| 285 |             | _ => error ("More than one type variable in type signature of constant " ^ quote c));
 | |
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 286 | val consts1 = | 
| 19038 | 287 | Locale.parameters_of thy name_locale | 
| 20892 | 288 | |> map (apsnd (Syntax.unlocalize_mixfix true)) | 
| 20175 | 289 | val SOME v = fold extract_classvar consts1 NONE; | 
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 290 | val consts2 = map ((apfst o apsnd) (subst_clsvar v (TFree (v, [])))) consts1; | 
| 19280 
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
 haftmann parents: 
19253diff
changeset | 291 | in (v, chop (length mapp_sup) consts2) end; | 
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 292 | fun add_consts v raw_cs_sup raw_cs_this thy = | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 293 | let | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 294 | fun add_global_const ((c, ty), syn) thy = | 
| 20265 | 295 | ((c, (Sign.full_name thy c, ty)), | 
| 296 | thy | |
| 20703 | 297 | |> Sign.add_consts_authentic [(c, ty |> subst_clsvar v (TFree (v, Sign.defaultS thy)), syn)]); | 
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 298 | in | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 299 | thy | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 300 | |> fold_map add_global_const raw_cs_this | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 301 | end; | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 302 | fun extract_assumes thy name_locale cs_mapp = | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 303 | let | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 304 | val subst_assume = | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 305 | map_aterms (fn Free (c, ty) => Const ((fst o the o AList.lookup (op =) cs_mapp) c, ty) | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 306 | | t => t) | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 307 | fun prep_asm ((name, atts), ts) = | 
| 19957 | 308 | ((NameSpace.base name, map (Attrib.attribute thy) atts), map subst_assume ts) | 
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 309 | in | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 310 | (map prep_asm o Locale.local_asms_of thy) name_locale | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 311 | end; | 
| 19280 
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
 haftmann parents: 
19253diff
changeset | 312 | fun add_global_constraint v class (_, (c, ty)) thy = | 
| 19038 | 313 | thy | 
| 19136 | 314 | |> Sign.add_const_constraint_i (c, SOME (subst_clsvar v (TFree (v, [class])) ty)); | 
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 315 | fun mk_const thy class v (c, ty) = | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 316 | Const (c, subst_clsvar v (TFree (v, [class])) ty); | 
| 20843 | 317 | fun is_operational thy mapp_this = | 
| 318 | length mapp_this > 0 orelse | |
| 319 | length (filter (#operational o fst o the o lookup_class_data thy) supclasses) > 1; | |
| 19038 | 320 | in | 
| 321 | thy | |
| 20428 | 322 | |> add_locale bname expr elems | 
| 20964 | 323 | |-> (fn name_locale => ProofContext.theory | 
| 324 | (`(fn thy => extract_tyvar_consts thy name_locale) | |
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 325 | #-> (fn (v, (raw_cs_sup, raw_cs_this)) => | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 326 | add_consts v raw_cs_sup raw_cs_this | 
| 19280 
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
 haftmann parents: 
19253diff
changeset | 327 | #-> (fn mapp_this => | 
| 
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
 haftmann parents: 
19253diff
changeset | 328 | `(fn thy => extract_assumes thy name_locale (mapp_sup @ mapp_this)) | 
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 329 | #-> (fn loc_axioms => | 
| 19957 | 330 | add_axclass_i (bname, supsort) (map (fst o snd) mapp_this) loc_axioms | 
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 331 | #-> (fn (name_axclass, (_, ax_axioms)) => | 
| 19280 
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
 haftmann parents: 
19253diff
changeset | 332 | fold (add_global_constraint v name_axclass) mapp_this | 
| 20843 | 333 | #> `(fn thy => is_operational thy mapp_this) | 
| 334 | #-> (fn operational => add_class_data (name_locale, (name_locale, name_axclass, v, mapp_this, | |
| 335 | operational, map (fst o fst) loc_axioms))) | |
| 20601 | 336 | #> prove_interpretation_i (bname, []) | 
| 19280 
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
 haftmann parents: 
19253diff
changeset | 337 | (Locale.Locale name_locale) (map (SOME o mk_const thy name_axclass v) (map snd (mapp_sup @ mapp_this))) | 
| 19929 | 338 | ((ALLGOALS o ProofContext.fact_tac) ax_axioms) | 
| 20964 | 339 | ))))) #> pair name_locale) | 
| 19038 | 340 | end; | 
| 341 | ||
| 342 | in | |
| 343 | ||
| 20385 | 344 | val class = gen_class (Locale.add_locale false) read_class; | 
| 19928 | 345 | val class_i = gen_class (Locale.add_locale_i false) certify_class; | 
| 19038 | 346 | |
| 20455 | 347 | end; (*local*) | 
| 19038 | 348 | |
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 349 | local | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 350 | |
| 20175 | 351 | fun gen_read_def thy prep_att read_def ((raw_name, raw_atts), raw_t) = | 
| 19038 | 352 | let | 
| 19957 | 353 | val (_, t) = read_def thy (raw_name, raw_t); | 
| 354 | val ((c, ty), _) = Sign.cert_def (Sign.pp thy) t; | |
| 355 | val atts = map (prep_att thy) raw_atts; | |
| 20175 | 356 | val insts = (Consts.typargs (Sign.consts_of thy) (c, ty)) | 
| 19957 | 357 | val name = case raw_name | 
| 20385 | 358 | of "" => NONE | 
| 359 | | _ => SOME raw_name; | |
| 20175 | 360 | in (c, (insts, ((name, t), atts))) end; | 
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 361 | |
| 19957 | 362 | fun read_def thy = gen_read_def thy Attrib.attribute read_axm; | 
| 363 | fun read_def_i thy = gen_read_def thy (K I) (K I); | |
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 364 | |
| 20175 | 365 | fun gen_instance_arity prep_arity prep_att read_def do_proof raw_arities (raw_name, raw_atts) raw_defs theory = | 
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 366 | let | 
| 20175 | 367 | fun prep_arity' ((tyco, asorts), sort) = prep_arity theory (tyco, asorts, sort); | 
| 368 | val arities = map prep_arity' raw_arities; | |
| 369 | val arities_pair = map (fn (tyco, asorts, sort) => ((tyco, asorts), sort)) arities; | |
| 370 | val _ = if null arities then error "at least one arity must be given" else (); | |
| 371 | val _ = case (duplicates (op =) o map #1) arities | |
| 372 | of [] => () | |
| 373 |       | dupl_tycos => error ("type constructors occur more than once in arities: "
 | |
| 374 | ^ (commas o map quote) dupl_tycos); | |
| 20393 | 375 | val (bind_always, name) = case raw_name | 
| 376 | of "" => (false, Thm.def_name ((space_implode "_" o map NameSpace.base) | |
| 20175 | 377 | (maps (fn (tyco, _, sort) => sort @ [tyco]) | 
| 20393 | 378 | (sort (fn ((tyco1, _, _), (tyco2, _, _)) => string_ord (tyco1, tyco2)) arities)))) | 
| 379 | | _ => (true, raw_name); | |
| 19136 | 380 | val atts = map (prep_att theory) raw_atts; | 
| 20175 | 381 | fun already_defined (c, ty_inst) = | 
| 382 |       is_some (find_first (fn (_, { lhs = [ty_inst'], ...}) =>
 | |
| 383 | Sign.typ_instance theory (ty_inst', ty_inst) orelse Sign.typ_instance theory (ty_inst, ty_inst')) | |
| 384 | (Defs.specifications_of (Theory.defs_of theory) c)); | |
| 385 | fun get_consts_class tyco ty class = | |
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 386 | let | 
| 19966 | 387 | val data = (fst o the_class_data theory) class; | 
| 19957 | 388 | val subst_ty = map_type_tfree (fn (v, sort) => | 
| 20175 | 389 | if #var data = v then ty else TVar ((v, 0), sort)); | 
| 19957 | 390 | in | 
| 391 | (map_filter (fn (_, (c, ty)) => | |
| 20175 | 392 | if already_defined (c, ty) | 
| 393 | then NONE else SOME ((c, ((tyco, class), subst_ty ty)))) o #consts) data | |
| 19957 | 394 | end; | 
| 20175 | 395 | fun get_consts_sort (tyco, asorts, sort) = | 
| 19957 | 396 | let | 
| 20192 
956cd30ef3be
renamed Name.give_names to Name.names and moved Name.alphanum to Symbol.alphanum
 haftmann parents: 
20191diff
changeset | 397 | val ty = Type (tyco, map (fn (v, sort) => TVar ((v, 0), sort)) (Name.names Name.context "'a" asorts)) | 
| 20175 | 398 | in maps (get_consts_class tyco ty) (the_ancestry theory sort) end; | 
| 399 | val cs = maps get_consts_sort arities; | |
| 20385 | 400 | fun mk_typnorm thy (ty, ty_sc) = | 
| 401 | case try (Sign.typ_match thy (Logic.varifyT ty_sc, ty)) Vartab.empty | |
| 402 | of SOME env => SOME (Logic.varifyT #> Envir.typ_subst_TVars env #> Logic.unvarifyT) | |
| 403 | | NONE => NONE; | |
| 20175 | 404 | fun read_defs defs cs thy_read = | 
| 405 | let | |
| 19957 | 406 | fun read raw_def cs = | 
| 407 | let | |
| 20385 | 408 | val (c, (inst, ((name_opt, t), atts))) = read_def thy_read raw_def; | 
| 409 | val ty = Consts.instance (Sign.consts_of thy_read) (c, inst); | |
| 20175 | 410 | val ((tyco, class), ty') = case AList.lookup (op =) cs c | 
| 19957 | 411 |              of NONE => error ("superfluous definition for constant " ^ quote c)
 | 
| 19966 | 412 | | SOME class_ty => class_ty; | 
| 20385 | 413 | val name = case name_opt | 
| 20628 | 414 | of NONE => Thm.def_name (Logic.name_arity (tyco, [], c)) | 
| 20385 | 415 | | SOME name => name; | 
| 416 | val t' = case mk_typnorm thy_read (ty', ty) | |
| 19966 | 417 |              of NONE => error ("superfluous definition for constant " ^
 | 
| 418 | quote c ^ "::" ^ Sign.string_of_typ thy_read ty) | |
| 20548 
8ef25fe585a8
renamed Term.map_term_types to Term.map_types (cf. Term.fold_types);
 wenzelm parents: 
20465diff
changeset | 419 | | SOME norm => map_types norm t | 
| 20175 | 420 | in (((class, tyco), ((name, t'), atts)), AList.delete (op =) c cs) end; | 
| 19957 | 421 | in fold_map read defs cs end; | 
| 20175 | 422 | val (defs, _) = assume_arities_thy theory arities_pair (read_defs raw_defs cs); | 
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 423 | fun get_remove_contraint c thy = | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 424 | let | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 425 | val ty = Sign.the_const_constraint thy c; | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 426 | in | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 427 | thy | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 428 | |> Sign.add_const_constraint_i (c, NONE) | 
| 20385 | 429 | |> pair (c, Logic.unvarifyT ty) | 
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 430 | end; | 
| 19966 | 431 | fun add_defs defs thy = | 
| 432 | thy | |
| 433 | |> PureThy.add_defs_i true (map snd defs) | |
| 434 | |-> (fn thms => pair (map fst defs ~~ thms)); | |
| 435 | fun note_all thy = | |
| 20175 | 436 | (*FIXME: should avoid binding duplicated names*) | 
| 19966 | 437 | let | 
| 20393 | 438 | val bind = bind_always orelse not (can (PureThy.get_thms thy) (Name name)); | 
| 20175 | 439 | val thms = maps (fn (tyco, _, sort) => maps (fn class => | 
| 440 | Symtab.lookup_list | |
| 441 | ((snd o the_class_data thy) class) tyco) (the_ancestry thy sort)) arities; | |
| 20393 | 442 | in if bind then | 
| 19966 | 443 | thy | 
| 20601 | 444 | |> PureThy.note_thmss_i (*qualified*) PureThy.internalK [((name, atts), [(thms, [])])] | 
| 19966 | 445 | |> snd | 
| 20393 | 446 | else | 
| 447 | thy | |
| 19966 | 448 | end; | 
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 449 | fun after_qed cs thy = | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 450 | thy | 
| 19412 | 451 | |> fold Sign.add_const_constraint_i (map (apsnd SOME) cs); | 
| 19038 | 452 | in | 
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 453 | theory | 
| 20385 | 454 | |> fold_map get_remove_contraint (map fst cs |> distinct (op =)) | 
| 19966 | 455 | ||>> add_defs defs | 
| 20076 | 456 | |-> (fn (cs, def_thms) => | 
| 20175 | 457 | fold add_inst_def def_thms | 
| 19966 | 458 | #> note_all | 
| 20182 
79c9ff40d760
tactic for prove_instance_arity now gets definition theorems as arguments
 haftmann parents: 
20175diff
changeset | 459 | #> do_proof (map snd def_thms) (after_qed cs) arities) | 
| 19038 | 460 | end; | 
| 461 | ||
| 19957 | 462 | fun instance_arity' do_proof = gen_instance_arity Sign.read_arity Attrib.attribute | 
| 463 | read_def do_proof; | |
| 464 | fun instance_arity_i' do_proof = gen_instance_arity Sign.cert_arity (K I) | |
| 465 | read_def_i do_proof; | |
| 20182 
79c9ff40d760
tactic for prove_instance_arity now gets definition theorems as arguments
 haftmann parents: 
20175diff
changeset | 466 | fun tactic_proof tac def_thms after_qed arities = | 
| 
79c9ff40d760
tactic for prove_instance_arity now gets definition theorems as arguments
 haftmann parents: 
20175diff
changeset | 467 | fold (fn arity => AxClass.prove_arity arity (tac def_thms)) arities | 
| 20175 | 468 | #> after_qed; | 
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 469 | |
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 470 | in | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 471 | |
| 20182 
79c9ff40d760
tactic for prove_instance_arity now gets definition theorems as arguments
 haftmann parents: 
20175diff
changeset | 472 | val instance_arity = instance_arity' (K axclass_instance_arity_i); | 
| 
79c9ff40d760
tactic for prove_instance_arity now gets definition theorems as arguments
 haftmann parents: 
20175diff
changeset | 473 | val instance_arity_i = instance_arity_i' (K axclass_instance_arity_i); | 
| 19150 | 474 | val prove_instance_arity = instance_arity_i' o tactic_proof; | 
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 475 | |
| 20455 | 476 | end; (*local*) | 
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 477 | |
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 478 | local | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 479 | |
| 20368 
2461a0485623
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 wenzelm parents: 
20353diff
changeset | 480 | fun prove_interpretation_in tac after_qed (name, expr) thy = | 
| 19136 | 481 | thy | 
| 20368 
2461a0485623
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 wenzelm parents: 
20353diff
changeset | 482 | |> Locale.interpretation_in_locale (ProofContext.theory after_qed) (name, expr) | 
| 19136 | 483 | |> Proof.global_terminal_proof (Method.Basic (fn _ => Method.SIMPLE_METHOD tac), NONE) | 
| 20368 
2461a0485623
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 wenzelm parents: 
20353diff
changeset | 484 | |> ProofContext.theory_of; | 
| 19038 | 485 | |
| 20385 | 486 | (*FIXME very ad-hoc, needs proper locale interface*) | 
| 19150 | 487 | fun gen_instance_sort prep_class prep_sort do_proof (raw_class, raw_sort) theory = | 
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 488 | let | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 489 | val class = prep_class theory raw_class; | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 490 | val sort = prep_sort theory raw_sort; | 
| 19966 | 491 | val loc_name = (#name_locale o fst o the_class_data theory) class; | 
| 19468 | 492 | val loc_expr = | 
| 19966 | 493 | (Locale.Merge o map (Locale.Locale o #name_locale o fst o the_class_data theory)) sort; | 
| 494 | val const_names = (map (NameSpace.base o fst o snd) | |
| 495 | o maps (#consts o fst o the_class_data theory) | |
| 496 | o the_ancestry theory) [class]; | |
| 20106 | 497 | fun get_thms thy = | 
| 498 | the_ancestry thy sort | |
| 499 | |> AList.make (the_propnames thy) | |
| 500 | |> map (apsnd (map (NameSpace.append (loc_name)))) | |
| 501 | |> map_filter (fn (superclass, thm_names) => | |
| 502 | case map_filter (try (PureThy.get_thm thy o Name)) thm_names | |
| 503 | of [] => NONE | |
| 504 | | thms => SOME (superclass, thms)); | |
| 505 | fun after_qed thy = | |
| 506 | thy | |
| 507 | |> `get_thms | |
| 508 | |-> fold (fn (supclass, thms) => I | |
| 509 | AxClass.prove_classrel (class, supclass) | |
| 510 | (ALLGOALS (K (intro_classes_tac [])) THEN | |
| 511 | (ALLGOALS o ProofContext.fact_tac) thms)) | |
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 512 | in | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 513 | theory | 
| 19136 | 514 | |> do_proof after_qed (loc_name, loc_expr) | 
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 515 | end; | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 516 | |
| 20385 | 517 | fun instance_sort' do_proof = gen_instance_sort read_class read_sort do_proof; | 
| 19150 | 518 | fun instance_sort_i' do_proof = gen_instance_sort certify_class certify_sort do_proof; | 
| 20368 
2461a0485623
global goals/qeds: after_qed operates on Proof.context (potentially local_theory);
 wenzelm parents: 
20353diff
changeset | 519 | val setup_proof = Locale.interpretation_in_locale o ProofContext.theory; | 
| 19136 | 520 | val tactic_proof = prove_interpretation_in; | 
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 521 | |
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 522 | in | 
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 523 | |
| 19150 | 524 | val instance_sort = instance_sort' setup_proof; | 
| 525 | val instance_sort_i = instance_sort_i' setup_proof; | |
| 526 | val prove_instance_sort = instance_sort_i' o tactic_proof; | |
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 527 | |
| 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 528 | end; (* local *) | 
| 19038 | 529 | |
| 19957 | 530 | |
| 18168 | 531 | |
| 20455 | 532 | (** code generation view **) | 
| 18168 | 533 | |
| 20843 | 534 | fun is_operational_class thy class = | 
| 535 | the_default false ((Option.map (#operational o fst) o lookup_class_data thy) class); | |
| 536 | ||
| 20455 | 537 | fun operational_algebra thy = | 
| 538 | Sorts.project_algebra (Sign.pp thy) | |
| 539 | (is_operational_class thy) (Sign.classes_of thy); | |
| 18168 | 540 | |
| 20455 | 541 | fun the_consts_sign thy class = | 
| 542 | let | |
| 543 |     val _ = if is_operational_class thy class then () else error ("no operational class: " ^ quote class);
 | |
| 544 | val data = (fst o the_class_data thy) class | |
| 545 | in (#var data, (map snd o #consts) data) end; | |
| 19213 | 546 | |
| 20455 | 547 | fun the_inst_sign thy (class, tyco) = | 
| 18168 | 548 | let | 
| 20455 | 549 |     val _ = if is_operational_class thy class then () else error ("no operational class: " ^ quote class);
 | 
| 550 | val asorts = Sign.arity_sorts thy tyco [class]; | |
| 551 | val (clsvar, const_sign) = the_consts_sign thy class; | |
| 552 | fun add_var sort used = | |
| 553 | let val v = hd (Name.invents used "'a" 1); | |
| 554 | in ((v, sort), Name.declare v used) end; | |
| 555 | val (vsorts, _) = | |
| 556 | Name.context | |
| 557 | |> Name.declare clsvar | |
| 558 | |> fold (fn (_, ty) => fold Name.declare | |
| 559 | ((map (fst o fst) o typ_tvars) ty @ map fst (Term.add_tfreesT ty []))) const_sign | |
| 560 | |> fold_map add_var asorts; | |
| 561 | val ty_inst = Type (tyco, map TFree vsorts); | |
| 562 | val inst_signs = map (apsnd (subst_clsvar clsvar ty_inst)) const_sign; | |
| 563 | in (vsorts, inst_signs) end; | |
| 18168 | 564 | |
| 18884 | 565 | |
| 20455 | 566 | |
| 567 | (** toplevel interface **) | |
| 18515 | 568 | |
| 569 | local | |
| 570 | ||
| 571 | structure P = OuterParse | |
| 572 | and K = OuterKeyword | |
| 573 | ||
| 574 | in | |
| 575 | ||
| 20385 | 576 | val (classK, instanceK, print_classesK) = ("class", "instance", "print_classes")
 | 
| 18515 | 577 | |
| 20385 | 578 | fun wrap_add_instance_sort (class, sort) thy = | 
| 579 | (if forall (is_some o lookup_class_data thy) (Sign.read_sort thy sort) | |
| 21094 | 580 | andalso (is_some o lookup_class_data thy) (Sign.intern_class thy class) | 
| 20385 | 581 | then instance_sort else axclass_instance_sort) (class, sort) thy; | 
| 19136 | 582 | |
| 20385 | 583 | val class_subP = P.name -- Scan.repeat (P.$$$ "+" |-- P.name) >> (op ::); | 
| 20428 | 584 | val class_bodyP = P.!!! (Scan.repeat1 P.locale_element); | 
| 20385 | 585 | |
| 20601 | 586 | val parse_arity = | 
| 587 | (P.xname --| P.$$$ "::" -- P.!!! P.arity) | |
| 19136 | 588 | >> (fn (tyco, (asorts, sort)) => ((tyco, asorts), sort)); | 
| 18911 | 589 | |
| 18515 | 590 | val classP = | 
| 20385 | 591 | OuterSyntax.command classK "define operational type class" K.thy_decl ( | 
| 18849 
05a16861d3a5
added three times overloaded Isar instance command
 haftmann parents: 
18829diff
changeset | 592 | P.name --| P.$$$ "=" | 
| 19280 
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
 haftmann parents: 
19253diff
changeset | 593 | -- ( | 
| 
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
 haftmann parents: 
19253diff
changeset | 594 | class_subP --| P.$$$ "+" -- class_bodyP | 
| 
5091dc43817b
slight improvement in serializer, stub for code generator theorems added
 haftmann parents: 
19253diff
changeset | 595 | || class_subP >> rpair [] | 
| 20964 | 596 | || class_bodyP >> pair []) | 
| 597 | -- P.opt_begin | |
| 598 | >> (fn ((bname, (supclasses, elems)), begin) => | |
| 20986 | 599 | Toplevel.begin_local_theory begin (class bname supclasses elems #-> TheoryTarget.begin))); | 
| 18515 | 600 | |
| 18575 | 601 | val instanceP = | 
| 18849 
05a16861d3a5
added three times overloaded Isar instance command
 haftmann parents: 
18829diff
changeset | 602 | OuterSyntax.command instanceK "prove type arity or subclass relation" K.thy_goal (( | 
| 20385 | 603 | P.xname -- ((P.$$$ "\\<subseteq>" || P.$$$ "<") |-- P.!!! P.xname) >> wrap_add_instance_sort | 
| 20601 | 604 | || P.opt_thm_name ":" -- (P.and_list1 parse_arity -- Scan.repeat (P.opt_thm_name ":" -- P.prop)) | 
| 20175 | 605 |            >> (fn (("", []), ([((tyco, asorts), sort)], [])) => axclass_instance_arity I [(tyco, asorts, sort)]
 | 
| 606 | | (natts, (arities, defs)) => instance_arity arities natts defs) | |
| 18849 
05a16861d3a5
added three times overloaded Isar instance command
 haftmann parents: 
18829diff
changeset | 607 | ) >> (Toplevel.print oo Toplevel.theory_to_proof)); | 
| 18575 | 608 | |
| 20385 | 609 | val print_classesP = | 
| 610 | OuterSyntax.improper_command print_classesK "print classes of this theory" K.diag | |
| 611 | (Scan.succeed (Toplevel.no_timing o Toplevel.unknown_theory | |
| 612 | o Toplevel.keep (print_classes o Toplevel.theory_of))); | |
| 613 | ||
| 614 | val _ = OuterSyntax.add_parsers [classP, instanceP, print_classesP]; | |
| 19110 
4bda27adcd2e
moved intro_classes from AxClass to ClassPackage
 haftmann parents: 
19102diff
changeset | 615 | |
| 20455 | 616 | end; (*local*) | 
| 18515 | 617 | |
| 20455 | 618 | end; |