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