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