src/Pure/axclass.ML
author berghofe
Fri, 24 Feb 2006 17:48:17 +0100
changeset 19134 47d337aefc21
parent 19123 a278d1e65c1d
child 19243 5dcb899a8486
permissions -rw-r--r--
Reverted to old interface of AxClass.add_inst_arity(_i)
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
     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
c20fbe3cb94f fixed types of add_XXX;
wenzelm
parents: 3854
diff changeset
     9
sig
5685
1e5b4c66317f quiet_mode, message;
wenzelm
parents: 4934
diff changeset
    10
  val quiet_mode: bool ref
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
    11
  val print_axclasses: theory -> unit
17281
3b9ff0131ed1 name space prefix is now "c_class" instead of just "c";
wenzelm
parents: 17221
diff changeset
    12
  val get_info: theory -> string -> {super_classes: class list, intro: thm, axioms: thm list}
17339
ab97ccef124a tuned Isar interfaces;
wenzelm
parents: 17281
diff changeset
    13
  val add_axclass: bstring * xstring list -> ((bstring * string) * Attrib.src list) list ->
18418
bf448d999b7e re-arranged tuples (theory * 'a) to ('a * theory) in Pure
haftmann
parents: 18377
diff changeset
    14
    theory -> {intro: thm, axioms: thm list} * theory
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
    15
  val add_axclass_i: bstring * class list -> ((bstring * term) * attribute list) list ->
18418
bf448d999b7e re-arranged tuples (theory * 'a) to ('a * theory) in Pure
haftmann
parents: 18377
diff changeset
    16
    theory -> {intro: thm, axioms: thm list} * theory
15876
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
    17
  val add_classrel_thms: thm list -> theory -> theory
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
    18
  val add_arity_thms: thm list -> theory -> theory
16333
490d77820631 got rid of bclass, xclass;
wenzelm
parents: 16180
diff changeset
    19
  val add_inst_subclass: xstring * xstring -> tactic -> theory -> theory
11828
ef3e51b1b4ec sane internal interfaces for instance;
wenzelm
parents: 11740
diff changeset
    20
  val add_inst_subclass_i: class * class -> tactic -> theory -> theory
19134
47d337aefc21 Reverted to old interface of AxClass.add_inst_arity(_i)
berghofe
parents: 19123
diff changeset
    21
  val add_inst_arity: xstring * string list * string -> tactic -> theory -> theory
47d337aefc21 Reverted to old interface of AxClass.add_inst_arity(_i)
berghofe
parents: 19123
diff changeset
    22
  val add_inst_arity_i: string * sort list * sort -> tactic -> theory -> theory
17339
ab97ccef124a tuned Isar interfaces;
wenzelm
parents: 17281
diff changeset
    23
  val instance_subclass: xstring * xstring -> theory -> Proof.state
ab97ccef124a tuned Isar interfaces;
wenzelm
parents: 17281
diff changeset
    24
  val instance_subclass_i: class * class -> theory -> Proof.state
18754
4e41252eae57 exported after_qed for axclass instance
haftmann
parents: 18728
diff changeset
    25
  val instance_arity: (theory -> theory) -> xstring * string list * string -> theory -> Proof.state
4e41252eae57 exported after_qed for axclass instance
haftmann
parents: 18728
diff changeset
    26
  val instance_arity_i: (theory -> theory) -> string * sort list * sort -> theory -> Proof.state
18574
46ed84a64cf6 exported read|cert_arity interfaces
haftmann
parents: 18467
diff changeset
    27
  val read_arity: theory -> xstring * string list * string -> arity
46ed84a64cf6 exported read|cert_arity interfaces
haftmann
parents: 18467
diff changeset
    28
  val cert_arity: theory -> string * sort list * sort -> arity
19123
a278d1e65c1d renamed class_axms to class_intros;
wenzelm
parents: 19110
diff changeset
    29
  val class_intros: theory -> thm list
3938
c20fbe3cb94f fixed types of add_XXX;
wenzelm
parents: 3854
diff changeset
    30
end;
404
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
    31
15801
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15705
diff changeset
    32
structure AxClass: AX_CLASS =
404
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
    33
struct
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
    34
4015
92874142156b add_store_axioms_i;
wenzelm
parents: 3988
diff changeset
    35
404
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
    36
(** utilities **)
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
    37
5685
1e5b4c66317f quiet_mode, message;
wenzelm
parents: 4934
diff changeset
    38
(* messages *)
1e5b4c66317f quiet_mode, message;
wenzelm
parents: 4934
diff changeset
    39
1e5b4c66317f quiet_mode, message;
wenzelm
parents: 4934
diff changeset
    40
val quiet_mode = ref false;
1e5b4c66317f quiet_mode, message;
wenzelm
parents: 4934
diff changeset
    41
fun message s = if ! quiet_mode then () else writeln s;
1e5b4c66317f quiet_mode, message;
wenzelm
parents: 4934
diff changeset
    42
1e5b4c66317f quiet_mode, message;
wenzelm
parents: 4934
diff changeset
    43
404
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
    44
(* type vars *)
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
    45
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
    46
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
    47
  | map_typ_frees f (TFree a) = f a
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
    48
  | map_typ_frees _ a = a;
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
val map_term_tfrees = map_term_types o map_typ_frees;
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 aT S = TFree ("'a", S);
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
    53
3395
d8700b008944 eliminated freeze_vars;
wenzelm
parents: 2961
diff changeset
    54
fun dest_varT (TFree (x, S)) = ((x, ~1), S)
d8700b008944 eliminated freeze_vars;
wenzelm
parents: 2961
diff changeset
    55
  | dest_varT (TVar xi_S) = xi_S
3788
51bd5c0954c6 eliminated raise_term, raise_typ;
wenzelm
parents: 3764
diff changeset
    56
  | dest_varT T = raise TYPE ("dest_varT", [T], []);
3395
d8700b008944 eliminated freeze_vars;
wenzelm
parents: 2961
diff changeset
    57
404
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
    58
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
    59
560
6702a715281d cleaned sig;
wenzelm
parents: 487
diff changeset
    60
(** abstract syntax operations **)
423
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
    61
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
    62
(* names *)
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
    63
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
    64
val introN = "intro";
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
    65
val axiomsN = "axioms";
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
    66
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
    67
423
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
    68
(* subclass relations as terms *)
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
    69
1498
7b7d20e89b1b Elimination of fully-functorial style.
paulson
parents: 1237
diff changeset
    70
fun mk_classrel (c1, c2) = Logic.mk_inclass (aT [c1], c2);
423
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
    71
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
    72
fun dest_classrel tm =
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
    73
  let
3788
51bd5c0954c6 eliminated raise_term, raise_typ;
wenzelm
parents: 3764
diff changeset
    74
    fun err () = raise TERM ("dest_classrel", [tm]);
423
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
    75
3395
d8700b008944 eliminated freeze_vars;
wenzelm
parents: 2961
diff changeset
    76
    val (ty, c2) = Logic.dest_inclass tm handle TERM _ => err ();
d8700b008944 eliminated freeze_vars;
wenzelm
parents: 2961
diff changeset
    77
    val c1 = (case dest_varT ty of (_, [c]) => c | _ => err ())
d8700b008944 eliminated freeze_vars;
wenzelm
parents: 2961
diff changeset
    78
      handle TYPE _ => err ();
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
    79
  in (c1, c2) end;
423
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
    80
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
    81
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
    82
(* arities as terms *)
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
    83
14605
9de4d64eee3b 'instance' and intro_classes now handle general sorts;
wenzelm
parents: 12876
diff changeset
    84
fun mk_arity (t, Ss, c) =
423
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
    85
  let
14695
9c78044b99c3 improved Term.invent_names;
wenzelm
parents: 14605
diff changeset
    86
    val tfrees = ListPair.map TFree (Term.invent_names [] "'a" (length Ss), Ss);
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
    87
  in Logic.mk_inclass (Type (t, tfrees), c) end;
423
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
    88
14605
9de4d64eee3b 'instance' and intro_classes now handle general sorts;
wenzelm
parents: 12876
diff changeset
    89
fun mk_arities (t, Ss, S) = map (fn c => mk_arity (t, Ss, c)) S;
9de4d64eee3b 'instance' and intro_classes now handle general sorts;
wenzelm
parents: 12876
diff changeset
    90
423
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
    91
fun dest_arity tm =
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
    92
  let
3788
51bd5c0954c6 eliminated raise_term, raise_typ;
wenzelm
parents: 3764
diff changeset
    93
    fun err () = raise TERM ("dest_arity", [tm]);
423
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
    94
3395
d8700b008944 eliminated freeze_vars;
wenzelm
parents: 2961
diff changeset
    95
    val (ty, c) = Logic.dest_inclass tm handle TERM _ => err ();
d8700b008944 eliminated freeze_vars;
wenzelm
parents: 2961
diff changeset
    96
    val (t, tvars) =
423
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
    97
      (case ty of
3395
d8700b008944 eliminated freeze_vars;
wenzelm
parents: 2961
diff changeset
    98
        Type (t, tys) => (t, map dest_varT tys handle TYPE _ => err ())
423
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
    99
      | _ => err ());
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
   100
    val ss =
18964
67f572e03236 renamed gen_duplicates to duplicates;
wenzelm
parents: 18932
diff changeset
   101
      if null (duplicates (eq_fst (op =)) tvars)
3395
d8700b008944 eliminated freeze_vars;
wenzelm
parents: 2961
diff changeset
   102
      then map snd tvars else err ();
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   103
  in (t, ss, c) end;
423
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
   104
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
   105
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
   106
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   107
(** axclass info **)
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   108
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   109
type axclass_info =
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   110
  {super_classes: class list,
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   111
    intro: thm,
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   112
    axioms: thm list};
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   113
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   114
structure AxclassesData = TheoryDataFun
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   115
(struct
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   116
  val name = "Pure/axclasses";
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   117
  type T = axclass_info Symtab.table;
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   118
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   119
  val empty = Symtab.empty;
6546
995a66249a9b theory data: copy;
wenzelm
parents: 6390
diff changeset
   120
  val copy = I;
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   121
  val extend = I;
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   122
  fun merge _ = Symtab.merge (K true);
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   123
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   124
  fun print thy tab =
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   125
    let
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   126
      fun pretty_class c cs = Pretty.block
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   127
        (Pretty.str (Sign.extern_class thy c) :: Pretty.str " <" :: Pretty.brk 1 ::
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   128
          Pretty.breaks (map (Pretty.str o Sign.extern_class thy) cs));
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   129
10008
61eb9f3aa92a Display.pretty_thm_sg;
wenzelm
parents: 8927
diff changeset
   130
      fun pretty_thms name thms = Pretty.big_list (name ^ ":")
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   131
        (map (Display.pretty_thm_sg thy) thms);
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   132
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   133
      fun pretty_axclass (name, {super_classes, intro, axioms}) = Pretty.block (Pretty.fbreaks
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   134
        [pretty_class name super_classes, pretty_thms introN [intro], pretty_thms axiomsN axioms]);
8720
840c75ab2a7f Pretty.chunks;
wenzelm
parents: 8716
diff changeset
   135
    in Pretty.writeln (Pretty.chunks (map pretty_axclass (Symtab.dest tab))) end;
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   136
end);
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   137
18708
4b3dadb4fe33 setup: theory -> theory;
wenzelm
parents: 18678
diff changeset
   138
val _ = Context.add_setup AxclassesData.init;
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   139
val print_axclasses = AxclassesData.print;
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   140
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   141
17412
e26cb20ef0cc TableFun/Symtab: curried lookup and update;
wenzelm
parents: 17339
diff changeset
   142
val lookup_info = Symtab.lookup o AxclassesData.get;
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   143
17281
3b9ff0131ed1 name space prefix is now "c_class" instead of just "c";
wenzelm
parents: 17221
diff changeset
   144
fun get_info thy c =
3b9ff0131ed1 name space prefix is now "c_class" instead of just "c";
wenzelm
parents: 17221
diff changeset
   145
  (case lookup_info thy c of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   146
    NONE => error ("Unknown axclass " ^ quote c)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   147
  | SOME info => info);
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   148
19123
a278d1e65c1d renamed class_axms to class_intros;
wenzelm
parents: 19110
diff changeset
   149
fun class_intros thy =
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   150
  let val classes = Sign.classes thy in
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   151
    map (Thm.class_triv thy) classes @
17281
3b9ff0131ed1 name space prefix is now "c_class" instead of just "c";
wenzelm
parents: 17221
diff changeset
   152
    List.mapPartial (Option.map #intro o lookup_info thy) classes
15876
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   153
  end;
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   154
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   155
423
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
   156
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
   157
(** add axiomatic type classes **)
404
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   158
15876
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   159
local
404
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   160
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   161
fun err_bad_axsort ax c =
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   162
  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
   163
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   164
fun err_bad_tfrees ax =
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   165
  error ("More than one type variable in axiom " ^ quote ax);
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   166
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   167
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
   168
  let
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   169
    val class = Sign.full_name thy bclass;
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   170
    val super_classes = map (prep_class thy) raw_super_classes;
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   171
    val axms = map (prep_axm thy o fst) raw_axioms_atts;
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   172
    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
   173
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   174
    (*declare class*)
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   175
    val class_thy =
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   176
      thy |> Theory.add_classes_i [(bclass, super_classes)];
404
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   177
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   178
    (*prepare abstract axioms*)
404
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   179
    fun abs_axm ax =
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   180
      if null (term_tfrees ax) then
14854
61bdf2ae4dc5 removed obsolete sort 'logic';
wenzelm
parents: 14824
diff changeset
   181
        Logic.mk_implies (Logic.mk_inclass (aT [], class), ax)
3788
51bd5c0954c6 eliminated raise_term, raise_typ;
wenzelm
parents: 3764
diff changeset
   182
      else map_term_tfrees (K (aT [class])) ax;
17756
d4a35f82fbb4 minor tweaks for Poplog/ML;
wenzelm
parents: 17703
diff changeset
   183
    val abs_axms = map (abs_axm o snd) axms;
404
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   184
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   185
    fun axm_sort (name, ax) =
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   186
      (case term_tfrees ax of
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   187
        [] => []
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   188
      | [(_, S)] => if Sign.subsort class_thy ([class], S) then S else err_bad_axsort name class
404
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   189
      | _ => err_bad_tfrees name);
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   190
    val axS = Sign.certify_sort class_thy (List.concat (map axm_sort axms));
404
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   191
1498
7b7d20e89b1b Elimination of fully-functorial style.
paulson
parents: 1237
diff changeset
   192
    val int_axm = Logic.close_form o map_term_tfrees (K (aT axS));
7b7d20e89b1b Elimination of fully-functorial style.
paulson
parents: 1237
diff changeset
   193
    fun inclass c = Logic.mk_inclass (aT axS, c);
404
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   194
1498
7b7d20e89b1b Elimination of fully-functorial style.
paulson
parents: 1237
diff changeset
   195
    val intro_axm = Logic.list_implies
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   196
      (map inclass super_classes @ map (int_axm o #2) axms, inclass class);
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   197
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   198
    (*declare axioms and rule*)
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18124
diff changeset
   199
    val (([intro], [axioms]), axms_thy) =
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   200
      class_thy
18932
66ecb05f92c8 Logic.const_of_class/class_of_const;
wenzelm
parents: 18848
diff changeset
   201
      |> Theory.add_path (Logic.const_of_class bclass)
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   202
      |> PureThy.add_axioms_i [Thm.no_attributes (introN, intro_axm)]
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18124
diff changeset
   203
      ||>> PureThy.add_axiomss_i [Thm.no_attributes (axiomsN, abs_axms)];
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   204
    val info = {super_classes = super_classes, intro = intro, axioms = axioms};
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   205
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   206
    (*store info*)
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18124
diff changeset
   207
    val (_, final_thy) =
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   208
      axms_thy
18932
66ecb05f92c8 Logic.const_of_class/class_of_const;
wenzelm
parents: 18848
diff changeset
   209
      |> Theory.add_finals_i false [Const (Logic.const_of_class class, a_itselfT --> propT)]
18377
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18124
diff changeset
   210
      |> PureThy.add_thms ((map #1 axms ~~ axioms) ~~ atts)
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18124
diff changeset
   211
      ||> Theory.restore_naming class_thy
0e1d025d57b3 oriented result pairs in PureThy
haftmann
parents: 18124
diff changeset
   212
      ||> AxclassesData.map (Symtab.update (class, info));
18418
bf448d999b7e re-arranged tuples (theory * 'a) to ('a * theory) in Pure
haftmann
parents: 18377
diff changeset
   213
  in ({intro = intro, axioms = axioms}, final_thy) end;
404
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   214
15876
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   215
in
404
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   216
18728
6790126ab5f6 simplified type attribute;
wenzelm
parents: 18708
diff changeset
   217
val add_axclass = ext_axclass Sign.intern_class Theory.read_axm Attrib.attribute;
6390
5d58c100ca3f qualify Theory.sign_of etc.;
wenzelm
parents: 6379
diff changeset
   218
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
   219
15876
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   220
end;
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   221
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   222
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   223
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   224
(** proven class instantiation **)
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   225
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   226
(* add thms to type signature *)
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   227
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   228
fun add_classrel_thms thms thy =
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   229
  let
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   230
    fun prep_thm thm =
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   231
      let
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   232
        val prop = Drule.plain_prop_of (Thm.transfer thy thm);
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   233
        val (c1, c2) = dest_classrel prop handle TERM _ =>
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   234
          raise THM ("add_classrel_thms: not a class relation", 0, [thm]);
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   235
      in (c1, c2) end;
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   236
  in Theory.add_classrel_i (map prep_thm thms) thy end;
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   237
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   238
fun add_arity_thms thms thy =
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   239
  let
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   240
    fun prep_thm thm =
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   241
      let
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   242
        val prop = Drule.plain_prop_of (Thm.transfer thy thm);
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   243
        val (t, ss, c) = dest_arity prop handle TERM _ =>
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   244
          raise THM ("add_arity_thms: not an arity", 0, [thm]);
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   245
      in (t, ss, [c]) end;
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   246
  in Theory.add_arities_i (map prep_thm thms) thy end;
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   247
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   248
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   249
(* prepare classes and arities *)
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   250
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   251
fun read_class thy c = Sign.certify_class thy (Sign.intern_class thy c);
15876
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   252
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   253
fun test_classrel thy cc = (Type.add_classrel (Sign.pp thy) [cc] (Sign.tsig_of thy); cc);
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   254
fun cert_classrel thy = test_classrel thy o Library.pairself (Sign.certify_class thy);
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   255
fun read_classrel thy = test_classrel thy o Library.pairself (read_class thy);
15876
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   256
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   257
fun test_arity thy ar = (Type.add_arities (Sign.pp thy) [ar] (Sign.tsig_of thy); ar);
15876
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   258
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   259
fun prep_arity prep_tycon prep_sort prep thy (t, Ss, x) =
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   260
  test_arity thy (prep_tycon thy t, map (prep_sort thy) Ss, prep thy x);
15876
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   261
16364
dc9f7066d80a refer to name spaces values instead of names;
wenzelm
parents: 16333
diff changeset
   262
val read_arity = prep_arity Sign.intern_type Sign.read_sort Sign.read_sort;
15876
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   263
val cert_arity = prep_arity (K I) Sign.certify_sort Sign.certify_sort;
404
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   264
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   265
15876
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   266
(* instance declarations -- tactical proof *)
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   267
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   268
local
423
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
   269
15876
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   270
fun ext_inst_subclass prep_classrel raw_cc tac thy =
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   271
  let
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   272
    val (c1, c2) = prep_classrel thy raw_cc;
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   273
    val txt = quote (Sign.string_of_classrel thy [c1, c2]);
15876
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   274
    val _ = message ("Proving class inclusion " ^ txt ^ " ...");
18678
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18574
diff changeset
   275
    val th = Goal.prove thy [] [] (mk_classrel (c1, c2)) (K tac) handle ERROR msg =>
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18574
diff changeset
   276
      cat_error msg ("The error(s) above occurred while trying to prove " ^ txt);
15876
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   277
  in add_classrel_thms [th] thy end;
404
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   278
19134
47d337aefc21 Reverted to old interface of AxClass.add_inst_arity(_i)
berghofe
parents: 19123
diff changeset
   279
fun ext_inst_arity prep_arity raw_arity tac thy =
15876
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   280
  let
16458
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   281
    val arity = prep_arity thy raw_arity;
4c6fd0c01d28 accomodate change of TheoryDataFun;
wenzelm
parents: 16364
diff changeset
   282
    val txt = quote (Sign.string_of_arity thy arity);
15876
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   283
    val _ = message ("Proving type arity " ^ txt ^ " ...");
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   284
    val props = (mk_arities arity);
17956
369e2af8ee45 Goal.prove;
wenzelm
parents: 17928
diff changeset
   285
    val ths = Goal.prove_multi thy [] [] props
18678
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18574
diff changeset
   286
      (fn _ => Tactic.precise_conjunction_tac (length props) 1 THEN tac) handle ERROR msg =>
dd0c569fa43d sane ERROR handling;
wenzelm
parents: 18574
diff changeset
   287
        cat_error msg ("\nThe error(s) above occurred while trying to prove " ^ txt);
19134
47d337aefc21 Reverted to old interface of AxClass.add_inst_arity(_i)
berghofe
parents: 19123
diff changeset
   288
  in add_arity_thms ths thy end;
15876
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   289
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   290
in
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   291
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   292
val add_inst_subclass = ext_inst_subclass read_classrel;
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   293
val add_inst_subclass_i = ext_inst_subclass cert_classrel;
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   294
val add_inst_arity = ext_inst_arity read_arity;
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   295
val add_inst_arity_i = ext_inst_arity cert_arity;
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   296
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   297
end;
404
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   298
423
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
   299
15876
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   300
(* instance declarations -- Isar proof *)
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   301
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   302
local
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   303
18754
4e41252eae57 exported after_qed for axclass instance
haftmann
parents: 18728
diff changeset
   304
fun gen_instance mk_prop add_thms after_qed inst thy =
4e41252eae57 exported after_qed for axclass instance
haftmann
parents: 18728
diff changeset
   305
  thy
17339
ab97ccef124a tuned Isar interfaces;
wenzelm
parents: 17281
diff changeset
   306
  |> ProofContext.init
18799
f137c5e971f5 moved theorem tags from Drule to PureThy;
wenzelm
parents: 18754
diff changeset
   307
  |> Proof.theorem_i PureThy.internalK NONE (after_qed oo fold add_thms) NONE ("", [])
18754
4e41252eae57 exported after_qed for axclass instance
haftmann
parents: 18728
diff changeset
   308
       (map (fn t => (("", []), [(t, ([], []))])) (mk_prop thy inst));
15876
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   309
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   310
in
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   311
17339
ab97ccef124a tuned Isar interfaces;
wenzelm
parents: 17281
diff changeset
   312
val instance_subclass =
18754
4e41252eae57 exported after_qed for axclass instance
haftmann
parents: 18728
diff changeset
   313
  gen_instance (single oo (mk_classrel oo read_classrel)) add_classrel_thms I;
17339
ab97ccef124a tuned Isar interfaces;
wenzelm
parents: 17281
diff changeset
   314
val instance_subclass_i =
18754
4e41252eae57 exported after_qed for axclass instance
haftmann
parents: 18728
diff changeset
   315
  gen_instance (single oo (mk_classrel oo cert_classrel)) add_classrel_thms I;
4e41252eae57 exported after_qed for axclass instance
haftmann
parents: 18728
diff changeset
   316
val instance_arity =
4e41252eae57 exported after_qed for axclass instance
haftmann
parents: 18728
diff changeset
   317
  gen_instance (mk_arities oo read_arity) add_arity_thms;
4e41252eae57 exported after_qed for axclass instance
haftmann
parents: 18728
diff changeset
   318
val instance_arity_i =
4e41252eae57 exported after_qed for axclass instance
haftmann
parents: 18728
diff changeset
   319
  gen_instance (mk_arities oo cert_arity) add_arity_thms;
15876
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   320
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   321
end;
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   322
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   323
423
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
   324
15876
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   325
(** outer syntax **)
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   326
17057
0934ac31985f OuterKeyword;
wenzelm
parents: 17041
diff changeset
   327
local structure P = OuterParse and K = OuterKeyword in
3949
c60ff6d0a6b8 fixed goal_XXX;
wenzelm
parents: 3938
diff changeset
   328
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   329
val axclassP =
6719
7c2dafc8e801 outer syntax keyword classification;
wenzelm
parents: 6679
diff changeset
   330
  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: 12694
diff changeset
   331
    ((P.name -- Scan.optional ((P.$$$ "\\<subseteq>" || P.$$$ "<") |--
a70df1e5bf10 got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents: 12694
diff changeset
   332
        P.!!! (P.list1 P.xname)) []) -- Scan.repeat P.spec_name
18418
bf448d999b7e re-arranged tuples (theory * 'a) to ('a * theory) in Pure
haftmann
parents: 18377
diff changeset
   333
      >> (Toplevel.theory o (snd oo uncurry add_axclass)));
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   334
18848
4ed69fe8f887 moved instance Isar command to class_package.ML
haftmann
parents: 18799
diff changeset
   335
val _ = OuterSyntax.add_parsers [axclassP];
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   336
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   337
end;
3949
c60ff6d0a6b8 fixed goal_XXX;
wenzelm
parents: 3938
diff changeset
   338
404
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   339
end;