| author | wenzelm | 
| Mon, 08 May 2000 11:45:57 +0200 | |
| changeset 8831 | b824c0c55613 | 
| parent 8720 | 840c75ab2a7f | 
| child 8897 | fb1436ca3b2e | 
| permissions | -rw-r--r-- | 
| 404 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 1 | (* Title: Pure/axclass.ML | 
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 2 | ID: $Id$ | 
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 3 | Author: Markus Wenzel, TU Muenchen | 
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 4 | |
| 6379 | 5 | Axiomatic type class package. | 
| 404 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 6 | *) | 
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 7 | |
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 8 | signature AX_CLASS = | 
| 3938 | 9 | sig | 
| 5685 | 10 | val quiet_mode: bool ref | 
| 6379 | 11 | val print_axclasses: theory -> unit | 
| 1498 | 12 | val add_classrel_thms: thm list -> theory -> theory | 
| 13 | val add_arity_thms: thm list -> theory -> theory | |
| 6379 | 14 | val add_axclass: bclass * xclass list -> ((bstring * string) * Args.src list) list | 
| 15 |     -> theory -> theory * {intro: thm, axioms: thm list}
 | |
| 16 | val add_axclass_i: bclass * class list -> ((bstring * term) * theory attribute list) list | |
| 17 |     -> theory -> theory * {intro: thm, axioms: thm list}
 | |
| 3938 | 18 | val add_inst_subclass: xclass * xclass -> string list -> thm list | 
| 19 | -> tactic option -> theory -> theory | |
| 3788 | 20 | val add_inst_subclass_i: class * class -> string list -> thm list | 
| 21 | -> tactic option -> theory -> theory | |
| 3938 | 22 | val add_inst_arity: xstring * xsort list * xclass list -> string list | 
| 23 | -> thm list -> tactic option -> theory -> theory | |
| 3788 | 24 | val add_inst_arity_i: string * sort list * class list -> string list | 
| 25 | -> thm list -> tactic option -> theory -> theory | |
| 6379 | 26 | val axclass_tac: thm list -> tactic | 
| 1498 | 27 | val prove_subclass: theory -> class * class -> thm list | 
| 28 | -> tactic option -> thm | |
| 29 | val prove_arity: theory -> string * sort list * class -> thm list | |
| 30 | -> tactic option -> thm | |
| 3949 | 31 | val goal_subclass: theory -> xclass * xclass -> thm list | 
| 32 | val goal_arity: theory -> xstring * xsort list * xclass -> thm list | |
| 6729 | 33 | val instance_subclass_proof: (xclass * xclass) * Comment.text -> bool -> theory -> ProofHistory.T | 
| 34 | val instance_subclass_proof_i: (class * class) * Comment.text -> bool -> theory -> ProofHistory.T | |
| 35 | val instance_arity_proof: (xstring * xsort list * xclass) * Comment.text | |
| 36 | -> bool -> theory -> ProofHistory.T | |
| 37 | val instance_arity_proof_i: (string * sort list * class) * Comment.text | |
| 38 | -> bool -> theory -> ProofHistory.T | |
| 6379 | 39 | val setup: (theory -> theory) list | 
| 3938 | 40 | end; | 
| 404 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 41 | |
| 1498 | 42 | structure AxClass : AX_CLASS = | 
| 404 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 43 | struct | 
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 44 | |
| 4015 | 45 | |
| 404 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 46 | (** utilities **) | 
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 47 | |
| 5685 | 48 | (* messages *) | 
| 49 | ||
| 50 | val quiet_mode = ref false; | |
| 51 | fun message s = if ! quiet_mode then () else writeln s; | |
| 52 | ||
| 53 | ||
| 404 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 54 | (* type vars *) | 
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 55 | |
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 56 | fun map_typ_frees f (Type (t, tys)) = Type (t, map (map_typ_frees f) tys) | 
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 57 | | map_typ_frees f (TFree a) = f a | 
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 58 | | map_typ_frees _ a = a; | 
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 59 | |
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 60 | val map_term_tfrees = map_term_types o map_typ_frees; | 
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 61 | |
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 62 | fun aT S = TFree ("'a", S);
 | 
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 63 | |
| 3395 | 64 | fun dest_varT (TFree (x, S)) = ((x, ~1), S) | 
| 65 | | dest_varT (TVar xi_S) = xi_S | |
| 3788 | 66 |   | dest_varT T = raise TYPE ("dest_varT", [T], []);
 | 
| 3395 | 67 | |
| 404 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 68 | |
| 886 
9af08725600b
instance: now automatically includes defs of current thy node as witnesses;
 wenzelm parents: 
638diff
changeset | 69 | (* get axioms and theorems *) | 
| 404 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 70 | |
| 1498 | 71 | val is_def = Logic.is_equals o #prop o rep_thm; | 
| 886 
9af08725600b
instance: now automatically includes defs of current thy node as witnesses;
 wenzelm parents: 
638diff
changeset | 72 | |
| 4934 | 73 | fun witnesses thy names thms = | 
| 6379 | 74 | PureThy.get_thmss thy names @ thms @ filter is_def (map snd (axioms_of thy)); | 
| 886 
9af08725600b
instance: now automatically includes defs of current thy node as witnesses;
 wenzelm parents: 
638diff
changeset | 75 | |
| 404 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 76 | |
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 77 | |
| 560 | 78 | (** abstract syntax operations **) | 
| 423 | 79 | |
| 6379 | 80 | (* names *) | 
| 81 | ||
| 82 | fun intro_name c = c ^ "I"; | |
| 83 | val introN = "intro"; | |
| 84 | val axiomsN = "axioms"; | |
| 85 | ||
| 86 | ||
| 423 | 87 | (* subclass relations as terms *) | 
| 88 | ||
| 1498 | 89 | fun mk_classrel (c1, c2) = Logic.mk_inclass (aT [c1], c2); | 
| 423 | 90 | |
| 91 | fun dest_classrel tm = | |
| 92 | let | |
| 3788 | 93 |     fun err () = raise TERM ("dest_classrel", [tm]);
 | 
| 423 | 94 | |
| 3395 | 95 | val (ty, c2) = Logic.dest_inclass tm handle TERM _ => err (); | 
| 96 | val c1 = (case dest_varT ty of (_, [c]) => c | _ => err ()) | |
| 97 | handle TYPE _ => err (); | |
| 6379 | 98 | in (c1, c2) end; | 
| 423 | 99 | |
| 100 | ||
| 101 | (* arities as terms *) | |
| 102 | ||
| 103 | fun mk_arity (t, ss, c) = | |
| 104 | let | |
| 449 | 105 | val names = tl (variantlist (replicate (length ss + 1) "'", [])); | 
| 2266 | 106 | val tfrees = ListPair.map TFree (names, ss); | 
| 6379 | 107 | in Logic.mk_inclass (Type (t, tfrees), c) end; | 
| 423 | 108 | |
| 109 | fun dest_arity tm = | |
| 110 | let | |
| 3788 | 111 |     fun err () = raise TERM ("dest_arity", [tm]);
 | 
| 423 | 112 | |
| 3395 | 113 | val (ty, c) = Logic.dest_inclass tm handle TERM _ => err (); | 
| 114 | val (t, tvars) = | |
| 423 | 115 | (case ty of | 
| 3395 | 116 | Type (t, tys) => (t, map dest_varT tys handle TYPE _ => err ()) | 
| 423 | 117 | | _ => err ()); | 
| 118 | val ss = | |
| 3395 | 119 | if null (gen_duplicates eq_fst tvars) | 
| 120 | then map snd tvars else err (); | |
| 6379 | 121 | in (t, ss, c) end; | 
| 423 | 122 | |
| 123 | ||
| 124 | ||
| 560 | 125 | (** add theorems as axioms **) | 
| 423 | 126 | |
| 127 | fun prep_thm_axm thy thm = | |
| 128 | let | |
| 129 |     fun err msg = raise THM ("prep_thm_axm: " ^ msg, 0, [thm]);
 | |
| 130 | ||
| 1237 | 131 |     val {sign, hyps, prop, ...} = rep_thm thm;
 | 
| 423 | 132 | in | 
| 6390 | 133 | if not (Sign.subsig (sign, Theory.sign_of thy)) then | 
| 423 | 134 | err "theorem not of same theory" | 
| 1237 | 135 | else if not (null (extra_shyps thm)) orelse not (null hyps) then | 
| 423 | 136 | err "theorem may not contain hypotheses" | 
| 137 | else prop | |
| 138 | end; | |
| 139 | ||
| 140 | (*theorems expressing class relations*) | |
| 141 | fun add_classrel_thms thms thy = | |
| 142 | let | |
| 143 | fun prep_thm thm = | |
| 144 | let | |
| 145 | val prop = prep_thm_axm thy thm; | |
| 146 | val (c1, c2) = dest_classrel prop handle TERM _ => | |
| 147 |           raise THM ("add_classrel_thms: theorem is not a class relation", 0, [thm]);
 | |
| 148 | in (c1, c2) end; | |
| 6379 | 149 | in Theory.add_classrel (map prep_thm thms) thy end; | 
| 423 | 150 | |
| 151 | (*theorems expressing arities*) | |
| 152 | fun add_arity_thms thms thy = | |
| 153 | let | |
| 154 | fun prep_thm thm = | |
| 155 | let | |
| 156 | val prop = prep_thm_axm thy thm; | |
| 157 | val (t, ss, c) = dest_arity prop handle TERM _ => | |
| 158 |           raise THM ("add_arity_thms: theorem is not an arity", 0, [thm]);
 | |
| 159 | in (t, ss, [c]) end; | |
| 6379 | 160 | in Theory.add_arities (map prep_thm thms) thy end; | 
| 161 | ||
| 162 | ||
| 163 | ||
| 164 | (** axclass info **) | |
| 165 | ||
| 166 | (* data kind 'Pure/axclasses' *) | |
| 167 | ||
| 168 | type axclass_info = | |
| 169 |   {super_classes: class list,
 | |
| 170 | intro: thm, | |
| 171 | axioms: thm list}; | |
| 172 | ||
| 173 | structure AxclassesArgs = | |
| 174 | struct | |
| 175 | val name = "Pure/axclasses"; | |
| 176 | type T = axclass_info Symtab.table; | |
| 177 | ||
| 178 | val empty = Symtab.empty; | |
| 6546 | 179 | val copy = I; | 
| 6379 | 180 | val prep_ext = I; | 
| 181 | fun merge (tab1, tab2) = Symtab.merge (K true) (tab1, tab2); | |
| 182 | ||
| 183 | fun print sg tab = | |
| 184 | let | |
| 185 | val ext_class = Sign.cond_extern sg Sign.classK; | |
| 186 | val ext_thm = PureThy.cond_extern_thm_sg sg; | |
| 187 | ||
| 188 | fun pretty_class c cs = Pretty.block | |
| 189 | (Pretty.str (ext_class c) :: Pretty.str " <" :: Pretty.brk 1 :: | |
| 190 | Pretty.breaks (map (Pretty.str o ext_class) cs)); | |
| 191 | ||
| 192 | fun pretty_thms name thms = Pretty.big_list (name ^ ":") (map Display.pretty_thm thms); | |
| 193 | ||
| 194 |       fun pretty_axclass (name, {super_classes, intro, axioms}) = Pretty.block (Pretty.fbreaks
 | |
| 195 | [pretty_class name super_classes, pretty_thms introN [intro], pretty_thms axiomsN axioms]); | |
| 8720 | 196 | in Pretty.writeln (Pretty.chunks (map pretty_axclass (Symtab.dest tab))) end; | 
| 6379 | 197 | end; | 
| 198 | ||
| 199 | structure AxclassesData = TheoryDataFun(AxclassesArgs); | |
| 200 | val print_axclasses = AxclassesData.print; | |
| 201 | ||
| 202 | ||
| 203 | (* get and put data *) | |
| 204 | ||
| 205 | fun lookup_axclass_info_sg sg c = Symtab.lookup (AxclassesData.get_sg sg, c); | |
| 206 | ||
| 207 | fun get_axclass_info thy c = | |
| 208 | (case lookup_axclass_info_sg (Theory.sign_of thy) c of | |
| 209 |     None => error ("Unknown axclass " ^ quote c)
 | |
| 210 | | Some info => info); | |
| 211 | ||
| 212 | fun put_axclass_info c info thy = | |
| 213 | thy |> AxclassesData.put (Symtab.update ((c, info), AxclassesData.get thy)); | |
| 423 | 214 | |
| 215 | ||
| 216 | ||
| 217 | (** add axiomatic type classes **) | |
| 404 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 218 | |
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 219 | (* errors *) | 
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 220 | |
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 221 | fun err_not_logic c = | 
| 4917 | 222 |   error ("Axiomatic class " ^ quote c ^ " not subclass of " ^ quote logicC);
 | 
| 404 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 223 | |
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 224 | fun err_bad_axsort ax c = | 
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 225 |   error ("Sort constraint in axiom " ^ quote ax ^ " not supersort of " ^ quote c);
 | 
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 226 | |
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 227 | fun err_bad_tfrees ax = | 
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 228 |   error ("More than one type variable in axiom " ^ quote ax);
 | 
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 229 | |
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 230 | |
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 231 | (* ext_axclass *) | 
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 232 | |
| 6379 | 233 | fun ext_axclass prep_class prep_axm prep_att (bclass, raw_super_classes) raw_axioms_atts thy = | 
| 404 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 234 | let | 
| 6379 | 235 | val sign = Theory.sign_of thy; | 
| 3938 | 236 | |
| 6379 | 237 | val class = Sign.full_name sign bclass; | 
| 238 | val super_classes = map (prep_class sign) raw_super_classes; | |
| 239 | val axms = map (prep_axm sign o fst) raw_axioms_atts; | |
| 240 | val atts = map (map (prep_att thy) o snd) raw_axioms_atts; | |
| 404 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 241 | |
| 6379 | 242 | (*declare class*) | 
| 243 | val class_thy = | |
| 244 | thy |> Theory.add_classes_i [(bclass, super_classes)]; | |
| 245 | val class_sign = Theory.sign_of class_thy; | |
| 404 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 246 | |
| 6379 | 247 | (*prepare abstract axioms*) | 
| 404 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 248 | fun abs_axm ax = | 
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 249 | if null (term_tfrees ax) then | 
| 1498 | 250 | Logic.mk_implies (Logic.mk_inclass (aT logicS, class), ax) | 
| 3788 | 251 | else map_term_tfrees (K (aT [class])) ax; | 
| 6379 | 252 | val abs_axms = map (abs_axm o #2) axms; | 
| 404 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 253 | |
| 6379 | 254 | (*prepare introduction rule*) | 
| 255 | val _ = if Sign.subsort class_sign ([class], logicS) then () else err_not_logic class; | |
| 404 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 256 | |
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 257 | fun axm_sort (name, ax) = | 
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 258 | (case term_tfrees ax of | 
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 259 | [] => [] | 
| 6379 | 260 | | [(_, S)] => if Sign.subsort class_sign ([class], S) then S else err_bad_axsort name class | 
| 404 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 261 | | _ => err_bad_tfrees name); | 
| 6379 | 262 | val axS = Sign.norm_sort class_sign (logicC :: flat (map axm_sort axms)); | 
| 404 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 263 | |
| 1498 | 264 | val int_axm = Logic.close_form o map_term_tfrees (K (aT axS)); | 
| 265 | fun inclass c = Logic.mk_inclass (aT axS, c); | |
| 404 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 266 | |
| 1498 | 267 | val intro_axm = Logic.list_implies | 
| 6379 | 268 | (map inclass super_classes @ map (int_axm o #2) axms, inclass class); | 
| 269 | ||
| 270 | (*declare axioms and rule*) | |
| 8420 | 271 | val (axms_thy, ([intro], [axioms])) = | 
| 6379 | 272 | class_thy | 
| 273 | |> Theory.add_path bclass | |
| 274 | |> PureThy.add_axioms_i [Thm.no_attributes (introN, intro_axm)] | |
| 8420 | 275 | |>>> PureThy.add_axiomss_i [Thm.no_attributes (axiomsN, abs_axms)]; | 
| 6379 | 276 |     val info = {super_classes = super_classes, intro = intro, axioms = axioms};
 | 
| 277 | ||
| 278 | (*store info*) | |
| 279 | val final_thy = | |
| 280 | axms_thy | |
| 8420 | 281 | |> (#1 o PureThy.add_thms ((map #1 axms ~~ axioms) ~~ atts)) | 
| 6379 | 282 | |> Theory.parent_path | 
| 8420 | 283 | |> (#1 o PureThy.add_thms [Thm.no_attributes (intro_name bclass, intro)]) | 
| 6379 | 284 | |> put_axclass_info class info; | 
| 285 |   in (final_thy, {intro = intro, axioms = axioms}) end;
 | |
| 404 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 286 | |
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 287 | |
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 288 | (* external interfaces *) | 
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 289 | |
| 6390 | 290 | val add_axclass = ext_axclass Sign.intern_class Theory.read_axm Attrib.global_attribute; | 
| 291 | val add_axclass_i = ext_axclass (K I) Theory.cert_axm (K I); | |
| 404 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 292 | |
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 293 | |
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 294 | |
| 423 | 295 | (** prove class relations and type arities **) | 
| 296 | ||
| 297 | (* class_axms *) | |
| 404 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 298 | |
| 6379 | 299 | fun class_axms sign = | 
| 300 | let val classes = Sign.classes sign in | |
| 301 | map (Thm.class_triv sign) classes @ | |
| 302 | mapfilter (apsome #intro o lookup_axclass_info_sg sign) classes | |
| 404 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 303 | end; | 
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 304 | |
| 423 | 305 | |
| 7351 | 306 | (* intro_classes *) | 
| 6379 | 307 | |
| 8520 | 308 | fun intro_classes_tac facts st = | 
| 8671 | 309 | HEADGOAL (Method.insert_tac facts THEN' | 
| 8520 | 310 | REPEAT_ALL_NEW (resolve_tac (class_axms (Thm.sign_of_thm st)))) st; | 
| 6379 | 311 | |
| 7351 | 312 | val intro_classes_method = | 
| 8520 | 313 |   ("intro_classes", Method.no_args (Method.METHOD intro_classes_tac),
 | 
| 7351 | 314 | "back-chain introduction rules of axiomatic type classes"); | 
| 6379 | 315 | |
| 316 | ||
| 423 | 317 | (* axclass_tac *) | 
| 318 | ||
| 487 
af83700cb771
added experimental add_defns (actually should be moved somewhere else);
 wenzelm parents: 
474diff
changeset | 319 | (*(1) repeatedly resolve goals of form "OFCLASS(ty, c_class)", | 
| 1217 | 320 | try class_trivs first, then "cI" axioms | 
| 423 | 321 | (2) rewrite goals using user supplied definitions | 
| 322 | (3) repeatedly resolve goals with user supplied non-definitions*) | |
| 323 | ||
| 6379 | 324 | fun axclass_tac thms = | 
| 1217 | 325 | let | 
| 326 | val defs = filter is_def thms; | |
| 327 | val non_defs = filter_out is_def thms; | |
| 328 | in | |
| 8520 | 329 | intro_classes_tac [] THEN | 
| 1217 | 330 | TRY (rewrite_goals_tac defs) THEN | 
| 331 | TRY (REPEAT_FIRST (fn i => assume_tac i ORELSE resolve_tac non_defs i)) | |
| 332 | end; | |
| 404 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 333 | |
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 334 | |
| 423 | 335 | (* provers *) | 
| 404 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 336 | |
| 6379 | 337 | fun prove mk_prop str_of thy sig_prop thms usr_tac = | 
| 404 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 338 | let | 
| 6390 | 339 | val sign = Theory.sign_of thy; | 
| 6679 | 340 | val goal = Thm.cterm_of sign (mk_prop sig_prop) | 
| 341 | handle TERM (msg, _) => error msg | |
| 342 | | TYPE (msg, _, _) => error msg; | |
| 6379 | 343 | val tac = axclass_tac thms THEN (if_none usr_tac all_tac); | 
| 423 | 344 | in | 
| 345 | prove_goalw_cterm [] goal (K [tac]) | |
| 346 | end | |
| 347 |   handle ERROR => error ("The error(s) above occurred while trying to prove "
 | |
| 6390 | 348 | ^ quote (str_of (Theory.sign_of thy, sig_prop))); | 
| 404 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 349 | |
| 638 | 350 | val prove_subclass = | 
| 3854 
762606a888fe
uses Sign.str_of_classrel, Sign.str_of_arity, Sign.str_of_arity;
 wenzelm parents: 
3809diff
changeset | 351 | prove mk_classrel (fn (sg, c1_c2) => Sign.str_of_classrel sg c1_c2); | 
| 404 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 352 | |
| 423 | 353 | val prove_arity = | 
| 3854 
762606a888fe
uses Sign.str_of_classrel, Sign.str_of_arity, Sign.str_of_arity;
 wenzelm parents: 
3809diff
changeset | 354 | prove mk_arity (fn (sg, (t, Ss, c)) => Sign.str_of_arity sg (t, Ss, [c])); | 
| 404 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 355 | |
| 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 356 | |
| 423 | 357 | |
| 449 | 358 | (** add proved subclass relations and arities **) | 
| 404 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 359 | |
| 3949 | 360 | fun intrn_classrel sg c1_c2 = | 
| 361 | pairself (Sign.intern_class sg) c1_c2; | |
| 362 | ||
| 8716 | 363 | fun intrn_arity intrn sg (raw_t, Ss, x) = | 
| 364 | let val t = Sign.intern_tycon sg raw_t in | |
| 365 |     if Sign.is_type_abbr sg t then error ("Illegal type abbreviation: " ^ quote t)
 | |
| 366 | else (t, map (Sign.intern_sort sg) Ss, intrn sg x) | |
| 367 | end; | |
| 6379 | 368 | |
| 369 | ||
| 370 | fun ext_inst_subclass prep_classrel raw_c1_c2 names thms usr_tac thy = | |
| 371 | let val c1_c2 = prep_classrel (Theory.sign_of thy) raw_c1_c2 in | |
| 5685 | 372 |     message ("Proving class inclusion " ^
 | 
| 6390 | 373 | quote (Sign.str_of_classrel (Theory.sign_of thy) c1_c2) ^ " ..."); | 
| 6379 | 374 | thy |> add_classrel_thms [prove_subclass thy c1_c2 (witnesses thy names thms) usr_tac] | 
| 3788 | 375 | end; | 
| 423 | 376 | |
| 6379 | 377 | fun ext_inst_arity prep_arity (raw_t, raw_Ss, raw_cs) names thms usr_tac thy = | 
| 423 | 378 | let | 
| 6379 | 379 | val sign = Theory.sign_of thy; | 
| 380 | val (t, Ss, cs) = prep_arity sign (raw_t, raw_Ss, raw_cs); | |
| 4934 | 381 | val wthms = witnesses thy names thms; | 
| 423 | 382 | fun prove c = | 
| 5685 | 383 |      (message ("Proving type arity " ^
 | 
| 3854 
762606a888fe
uses Sign.str_of_classrel, Sign.str_of_arity, Sign.str_of_arity;
 wenzelm parents: 
3809diff
changeset | 384 | quote (Sign.str_of_arity sign (t, Ss, [c])) ^ " ..."); | 
| 
762606a888fe
uses Sign.str_of_classrel, Sign.str_of_arity, Sign.str_of_arity;
 wenzelm parents: 
3809diff
changeset | 385 | prove_arity thy (t, Ss, c) wthms usr_tac); | 
| 6379 | 386 | in add_arity_thms (map prove cs) thy end; | 
| 387 | ||
| 388 | ||
| 389 | val add_inst_subclass = ext_inst_subclass intrn_classrel; | |
| 390 | val add_inst_subclass_i = ext_inst_subclass (K I); | |
| 391 | val add_inst_arity = ext_inst_arity (intrn_arity Sign.intern_sort); | |
| 392 | val add_inst_arity_i = ext_inst_arity (K I); | |
| 393 | ||
| 404 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 394 | |
| 6379 | 395 | (* make old-style interactive goals *) | 
| 396 | ||
| 397 | fun mk_goal mk_prop thy sig_prop = | |
| 398 | Goals.goalw_cterm [] (Thm.cterm_of (Theory.sign_of thy) (mk_prop (Theory.sign_of thy) sig_prop)); | |
| 399 | ||
| 400 | val goal_subclass = mk_goal (mk_classrel oo intrn_classrel); | |
| 401 | val goal_arity = mk_goal (mk_arity oo intrn_arity Sign.intern_class); | |
| 402 | ||
| 3788 | 403 | |
| 404 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 404 | |
| 6379 | 405 | (** instance proof interfaces **) | 
| 406 | ||
| 407 | fun inst_attribute add_thms (thy, thm) = (add_thms [thm] thy, thm); | |
| 408 | ||
| 6729 | 409 | fun inst_proof mk_prop add_thms (sig_prop, comment) int thy = | 
| 6379 | 410 | thy | 
| 6729 | 411 |   |> IsarThy.theorem_i (("", [inst_attribute add_thms],
 | 
| 6935 | 412 | (mk_prop (Theory.sign_of thy) sig_prop, ([], []))), comment) int; | 
| 3949 | 413 | |
| 6379 | 414 | val instance_subclass_proof = inst_proof (mk_classrel oo intrn_classrel) add_classrel_thms; | 
| 415 | val instance_subclass_proof_i = inst_proof (K mk_classrel) add_classrel_thms; | |
| 416 | val instance_arity_proof = inst_proof (mk_arity oo intrn_arity Sign.intern_class) add_arity_thms; | |
| 417 | val instance_arity_proof_i = inst_proof (K mk_arity) add_arity_thms; | |
| 418 | ||
| 419 | ||
| 420 | ||
| 421 | (** package setup **) | |
| 422 | ||
| 423 | (* setup theory *) | |
| 3949 | 424 | |
| 6379 | 425 | val setup = | 
| 426 | [AxclassesData.init, | |
| 7351 | 427 | Method.add_methods [intro_classes_method]]; | 
| 6379 | 428 | |
| 429 | ||
| 430 | (* outer syntax *) | |
| 431 | ||
| 6719 | 432 | local structure P = OuterParse and K = OuterSyntax.Keyword in | 
| 3949 | 433 | |
| 6379 | 434 | val axclassP = | 
| 6719 | 435 | OuterSyntax.command "axclass" "define axiomatic type class" K.thy_decl | 
| 6729 | 436 | (((P.name -- Scan.optional (P.$$$ "<" |-- P.!!! (P.list1 P.xname)) []) --| P.marg_comment) | 
| 437 | -- Scan.repeat (P.spec_name --| P.marg_comment) | |
| 6379 | 438 | >> (fn (cls, axs) => Toplevel.theory (#1 o add_axclass cls axs))); | 
| 439 | ||
| 440 | val instanceP = | |
| 6719 | 441 | OuterSyntax.command "instance" "prove type arity or subclass relation" K.thy_goal | 
| 6729 | 442 | ((P.xname -- (P.$$$ "<" |-- P.xname) -- P.marg_comment >> instance_subclass_proof || | 
| 443 | (P.xname -- (P.$$$ "::" |-- P.simple_arity) >> P.triple2) -- P.marg_comment | |
| 444 | >> instance_arity_proof) | |
| 6379 | 445 | >> (Toplevel.print oo Toplevel.theory_to_proof)); | 
| 446 | ||
| 447 | val _ = OuterSyntax.add_parsers [axclassP, instanceP]; | |
| 448 | ||
| 449 | end; | |
| 3949 | 450 | |
| 451 | ||
| 404 
dd3d3d6467db
axiomatic type class 'package' for Pure (alpha version);
 wenzelm parents: diff
changeset | 452 | end; |