src/Pure/axclass.ML
author kleing
Fri, 27 May 2005 01:09:44 +0200
changeset 16095 f6af6b265d20
parent 15973 5fd94d84470f
child 16122 864fda4a4056
permissions -rw-r--r--
put global isatest settings in one file, sourced by the other scripts
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
15703
727ef1b8b3ee *** empty log message ***
wenzelm
parents: 15696
diff changeset
    12
  val add_axclass: bclass * xclass list -> ((bstring * string) * Attrib.src list) list
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
    13
    -> theory -> theory * {intro: thm, axioms: thm list}
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
    14
  val add_axclass_i: bclass * class list -> ((bstring * term) * theory attribute list) list
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
    15
    -> theory -> theory * {intro: thm, axioms: thm list}
15876
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
    16
  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
    17
  val add_arity_thms: thm list -> theory -> theory
11828
ef3e51b1b4ec sane internal interfaces for instance;
wenzelm
parents: 11740
diff changeset
    18
  val add_inst_subclass: xclass * xclass -> tactic -> theory -> theory
ef3e51b1b4ec sane internal interfaces for instance;
wenzelm
parents: 11740
diff changeset
    19
  val add_inst_subclass_i: class * class -> tactic -> theory -> theory
ef3e51b1b4ec sane internal interfaces for instance;
wenzelm
parents: 11740
diff changeset
    20
  val add_inst_arity: xstring * string list * string -> tactic -> theory -> theory
ef3e51b1b4ec sane internal interfaces for instance;
wenzelm
parents: 11740
diff changeset
    21
  val add_inst_arity_i: string * sort list * sort -> tactic -> theory -> theory
12876
a70df1e5bf10 got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents: 12694
diff changeset
    22
  val instance_subclass_proof: xclass * xclass -> bool -> theory -> ProofHistory.T
a70df1e5bf10 got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents: 12694
diff changeset
    23
  val instance_subclass_proof_i: class * class -> bool -> theory -> ProofHistory.T
14605
9de4d64eee3b 'instance' and intro_classes now handle general sorts;
wenzelm
parents: 12876
diff changeset
    24
  val instance_arity_proof: xstring * string list * string -> bool -> theory -> ProofHistory.T
9de4d64eee3b 'instance' and intro_classes now handle general sorts;
wenzelm
parents: 12876
diff changeset
    25
  val instance_arity_proof_i: string * sort list * sort -> bool -> theory -> ProofHistory.T
15876
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
    26
  val intro_classes_tac: thm list -> tactic
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
    27
  val default_intro_classes_tac: thm list -> tactic
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
    28
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
    29
  (*legacy interfaces*)
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
    30
  val axclass_tac: thm list -> tactic
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
    31
  val add_inst_subclass_x: xclass * xclass -> string list -> thm list
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
    32
    -> tactic option -> 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
    33
  val add_inst_arity_x: xstring * string list * string -> string list
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
    34
    -> thm list -> tactic option -> theory -> theory
3938
c20fbe3cb94f fixed types of add_XXX;
wenzelm
parents: 3854
diff changeset
    35
end;
404
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
    36
15801
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15705
diff changeset
    37
structure AxClass: AX_CLASS =
404
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
    38
struct
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
    39
4015
92874142156b add_store_axioms_i;
wenzelm
parents: 3988
diff changeset
    40
404
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
    41
(** utilities **)
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
    42
5685
1e5b4c66317f quiet_mode, message;
wenzelm
parents: 4934
diff changeset
    43
(* messages *)
1e5b4c66317f quiet_mode, message;
wenzelm
parents: 4934
diff changeset
    44
1e5b4c66317f quiet_mode, message;
wenzelm
parents: 4934
diff changeset
    45
val quiet_mode = ref false;
1e5b4c66317f quiet_mode, message;
wenzelm
parents: 4934
diff changeset
    46
fun message s = if ! quiet_mode then () else writeln s;
1e5b4c66317f quiet_mode, message;
wenzelm
parents: 4934
diff changeset
    47
1e5b4c66317f quiet_mode, message;
wenzelm
parents: 4934
diff changeset
    48
404
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
    49
(* type vars *)
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
    50
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
    51
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
    52
  | map_typ_frees f (TFree a) = f a
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
    53
  | map_typ_frees _ a = a;
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
    54
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
    55
val map_term_tfrees = map_term_types o map_typ_frees;
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
    56
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
    57
fun aT S = TFree ("'a", S);
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
    58
3395
d8700b008944 eliminated freeze_vars;
wenzelm
parents: 2961
diff changeset
    59
fun dest_varT (TFree (x, S)) = ((x, ~1), S)
d8700b008944 eliminated freeze_vars;
wenzelm
parents: 2961
diff changeset
    60
  | dest_varT (TVar xi_S) = xi_S
3788
51bd5c0954c6 eliminated raise_term, raise_typ;
wenzelm
parents: 3764
diff changeset
    61
  | dest_varT T = raise TYPE ("dest_varT", [T], []);
3395
d8700b008944 eliminated freeze_vars;
wenzelm
parents: 2961
diff changeset
    62
404
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
    63
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
    64
560
6702a715281d cleaned sig;
wenzelm
parents: 487
diff changeset
    65
(** abstract syntax operations **)
423
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
    66
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
    67
(* names *)
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
    68
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
    69
fun intro_name c = c ^ "I";
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
    70
val introN = "intro";
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
    71
val axiomsN = "axioms";
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
    72
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
    73
423
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
    74
(* subclass relations as terms *)
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
    75
1498
7b7d20e89b1b Elimination of fully-functorial style.
paulson
parents: 1237
diff changeset
    76
fun mk_classrel (c1, c2) = Logic.mk_inclass (aT [c1], c2);
423
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
    77
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
    78
fun dest_classrel tm =
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
    79
  let
3788
51bd5c0954c6 eliminated raise_term, raise_typ;
wenzelm
parents: 3764
diff changeset
    80
    fun err () = raise TERM ("dest_classrel", [tm]);
423
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
    81
3395
d8700b008944 eliminated freeze_vars;
wenzelm
parents: 2961
diff changeset
    82
    val (ty, c2) = Logic.dest_inclass tm handle TERM _ => err ();
d8700b008944 eliminated freeze_vars;
wenzelm
parents: 2961
diff changeset
    83
    val c1 = (case dest_varT ty of (_, [c]) => c | _ => err ())
d8700b008944 eliminated freeze_vars;
wenzelm
parents: 2961
diff changeset
    84
      handle TYPE _ => err ();
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
    85
  in (c1, c2) end;
423
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
    86
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
    87
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
    88
(* arities as terms *)
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
    89
14605
9de4d64eee3b 'instance' and intro_classes now handle general sorts;
wenzelm
parents: 12876
diff changeset
    90
fun mk_arity (t, Ss, c) =
423
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
    91
  let
14695
9c78044b99c3 improved Term.invent_names;
wenzelm
parents: 14605
diff changeset
    92
    val tfrees = ListPair.map TFree (Term.invent_names [] "'a" (length Ss), Ss);
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
    93
  in Logic.mk_inclass (Type (t, tfrees), c) end;
423
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
    94
14605
9de4d64eee3b 'instance' and intro_classes now handle general sorts;
wenzelm
parents: 12876
diff changeset
    95
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
    96
423
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
    97
fun dest_arity tm =
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
    98
  let
3788
51bd5c0954c6 eliminated raise_term, raise_typ;
wenzelm
parents: 3764
diff changeset
    99
    fun err () = raise TERM ("dest_arity", [tm]);
423
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
   100
3395
d8700b008944 eliminated freeze_vars;
wenzelm
parents: 2961
diff changeset
   101
    val (ty, c) = Logic.dest_inclass tm handle TERM _ => err ();
d8700b008944 eliminated freeze_vars;
wenzelm
parents: 2961
diff changeset
   102
    val (t, tvars) =
423
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
   103
      (case ty of
3395
d8700b008944 eliminated freeze_vars;
wenzelm
parents: 2961
diff changeset
   104
        Type (t, tys) => (t, map dest_varT tys handle TYPE _ => err ())
423
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
   105
      | _ => err ());
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
   106
    val ss =
3395
d8700b008944 eliminated freeze_vars;
wenzelm
parents: 2961
diff changeset
   107
      if null (gen_duplicates eq_fst tvars)
d8700b008944 eliminated freeze_vars;
wenzelm
parents: 2961
diff changeset
   108
      then map snd tvars else err ();
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   109
  in (t, ss, c) end;
423
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
   110
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
   111
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
   112
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   113
(** axclass info **)
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   114
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   115
(* data kind 'Pure/axclasses' *)
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   116
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   117
type axclass_info =
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   118
  {super_classes: class list,
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   119
    intro: thm,
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   120
    axioms: thm list};
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   121
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   122
structure AxclassesArgs =
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   123
struct
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   124
  val name = "Pure/axclasses";
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   125
  type T = axclass_info Symtab.table;
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   126
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   127
  val empty = Symtab.empty;
6546
995a66249a9b theory data: copy;
wenzelm
parents: 6390
diff changeset
   128
  val copy = I;
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   129
  val prep_ext = I;
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   130
  fun merge (tab1, tab2) = Symtab.merge (K true) (tab1, tab2);
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   131
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   132
  fun print sg tab =
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   133
    let
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   134
      val ext_class = Sign.cond_extern sg Sign.classK;
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   135
      val ext_thm = PureThy.cond_extern_thm_sg sg;
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   136
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   137
      fun pretty_class c cs = Pretty.block
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   138
        (Pretty.str (ext_class c) :: Pretty.str " <" :: Pretty.brk 1 ::
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   139
          Pretty.breaks (map (Pretty.str o ext_class) cs));
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   140
10008
61eb9f3aa92a Display.pretty_thm_sg;
wenzelm
parents: 8927
diff changeset
   141
      fun pretty_thms name thms = Pretty.big_list (name ^ ":")
61eb9f3aa92a Display.pretty_thm_sg;
wenzelm
parents: 8927
diff changeset
   142
        (map (Display.pretty_thm_sg sg) thms);
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   143
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   144
      fun pretty_axclass (name, {super_classes, intro, axioms}) = Pretty.block (Pretty.fbreaks
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   145
        [pretty_class name super_classes, pretty_thms introN [intro], pretty_thms axiomsN axioms]);
8720
840c75ab2a7f Pretty.chunks;
wenzelm
parents: 8716
diff changeset
   146
    in Pretty.writeln (Pretty.chunks (map pretty_axclass (Symtab.dest tab))) end;
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   147
end;
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   148
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   149
structure AxclassesData = TheoryDataFun(AxclassesArgs);
15801
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15705
diff changeset
   150
val _ = Context.add_setup [AxclassesData.init];
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   151
val print_axclasses = AxclassesData.print;
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   152
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   153
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   154
(* get and put data *)
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   155
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   156
fun lookup_axclass_info_sg sg c = Symtab.lookup (AxclassesData.get_sg sg, c);
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   157
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   158
fun get_axclass_info thy c =
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   159
  (case lookup_axclass_info_sg (Theory.sign_of thy) c of
15531
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   160
    NONE => error ("Unknown axclass " ^ quote c)
08c8dad8e399 Deleted Library.option type.
skalberg
parents: 15457
diff changeset
   161
  | SOME info => info);
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   162
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   163
fun put_axclass_info c info thy =
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   164
  thy |> AxclassesData.put (Symtab.update ((c, info), AxclassesData.get thy));
423
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
   165
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
   166
15876
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   167
(* class_axms *)
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   168
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   169
fun class_axms sign =
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   170
  let val classes = Sign.classes sign in
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   171
    map (Thm.class_triv sign) classes @
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   172
    List.mapPartial (Option.map #intro o lookup_axclass_info_sg sign) classes
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   173
  end;
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   174
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   175
423
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
   176
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
   177
(** add axiomatic type classes **)
404
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   178
15876
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   179
local
404
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   180
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   181
fun err_bad_axsort ax c =
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   182
  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
   183
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   184
fun err_bad_tfrees ax =
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   185
  error ("More than one type variable in axiom " ^ quote ax);
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   186
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   187
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
   188
  let
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   189
    val sign = Theory.sign_of thy;
3938
c20fbe3cb94f fixed types of add_XXX;
wenzelm
parents: 3854
diff changeset
   190
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   191
    val class = Sign.full_name sign bclass;
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   192
    val super_classes = map (prep_class sign) raw_super_classes;
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   193
    val axms = map (prep_axm sign o fst) raw_axioms_atts;
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   194
    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
   195
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   196
    (*declare class*)
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   197
    val class_thy =
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   198
      thy |> Theory.add_classes_i [(bclass, super_classes)];
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   199
    val class_sign = Theory.sign_of class_thy;
404
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   200
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   201
    (*prepare abstract axioms*)
404
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   202
    fun abs_axm ax =
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   203
      if null (term_tfrees ax) then
14854
61bdf2ae4dc5 removed obsolete sort 'logic';
wenzelm
parents: 14824
diff changeset
   204
        Logic.mk_implies (Logic.mk_inclass (aT [], class), ax)
3788
51bd5c0954c6 eliminated raise_term, raise_typ;
wenzelm
parents: 3764
diff changeset
   205
      else map_term_tfrees (K (aT [class])) ax;
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   206
    val abs_axms = map (abs_axm o #2) axms;
404
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   207
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   208
    fun axm_sort (name, ax) =
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   209
      (case term_tfrees ax of
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   210
        [] => []
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   211
      | [(_, 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
   212
      | _ => err_bad_tfrees name);
15570
8d8c70b41bab Move towards standard functions.
skalberg
parents: 15531
diff changeset
   213
    val axS = Sign.certify_sort class_sign (List.concat (map axm_sort axms));
404
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   214
1498
7b7d20e89b1b Elimination of fully-functorial style.
paulson
parents: 1237
diff changeset
   215
    val int_axm = Logic.close_form o map_term_tfrees (K (aT axS));
7b7d20e89b1b Elimination of fully-functorial style.
paulson
parents: 1237
diff changeset
   216
    fun inclass c = Logic.mk_inclass (aT axS, c);
404
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   217
1498
7b7d20e89b1b Elimination of fully-functorial style.
paulson
parents: 1237
diff changeset
   218
    val intro_axm = Logic.list_implies
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   219
      (map inclass super_classes @ map (int_axm o #2) axms, inclass class);
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   220
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   221
    (*declare axioms and rule*)
8420
f37fd19476ca adapted to new PureThy.add_thms etc.;
wenzelm
parents: 7351
diff changeset
   222
    val (axms_thy, ([intro], [axioms])) =
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   223
      class_thy
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   224
      |> Theory.add_path bclass
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   225
      |> PureThy.add_axioms_i [Thm.no_attributes (introN, intro_axm)]
8420
f37fd19476ca adapted to new PureThy.add_thms etc.;
wenzelm
parents: 7351
diff changeset
   226
      |>>> PureThy.add_axiomss_i [Thm.no_attributes (axiomsN, abs_axms)];
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   227
    val info = {super_classes = super_classes, intro = intro, axioms = axioms};
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   228
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   229
    (*store info*)
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   230
    val final_thy =
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   231
      axms_thy
8420
f37fd19476ca adapted to new PureThy.add_thms etc.;
wenzelm
parents: 7351
diff changeset
   232
      |> (#1 o PureThy.add_thms ((map #1 axms ~~ axioms) ~~ atts))
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   233
      |> Theory.parent_path
8420
f37fd19476ca adapted to new PureThy.add_thms etc.;
wenzelm
parents: 7351
diff changeset
   234
      |> (#1 o PureThy.add_thms [Thm.no_attributes (intro_name bclass, intro)])
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   235
      |> put_axclass_info class info;
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   236
  in (final_thy, {intro = intro, axioms = axioms}) end;
404
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   237
15876
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   238
in
404
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   239
6390
5d58c100ca3f qualify Theory.sign_of etc.;
wenzelm
parents: 6379
diff changeset
   240
val add_axclass = ext_axclass Sign.intern_class Theory.read_axm Attrib.global_attribute;
5d58c100ca3f qualify Theory.sign_of etc.;
wenzelm
parents: 6379
diff changeset
   241
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
   242
15876
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   243
end;
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   244
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   245
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   246
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   247
(** 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
   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
(* 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
   250
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   251
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
   252
  let
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   253
    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
   254
      let
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   255
        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
   256
        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
   257
          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
   258
      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
   259
  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
   260
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   261
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
   262
  let
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   263
    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
   264
      let
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   265
        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
   266
        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
   267
          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
   268
      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
   269
  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
   270
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   271
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   272
(* 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
   273
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   274
fun read_class sg c = Sign.certify_class sg (Sign.intern_class sg c);
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   275
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   276
fun test_classrel sg cc = (Type.add_classrel (Sign.pp sg) [cc] (Sign.tsig_of sg); cc);
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   277
fun cert_classrel sg = test_classrel sg o Library.pairself (Sign.certify_class sg);
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   278
fun read_classrel sg = test_classrel sg o Library.pairself (read_class sg);
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   279
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   280
fun test_arity sg ar = (Type.add_arities (Sign.pp sg) [ar] (Sign.tsig_of sg); ar);
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   281
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   282
fun prep_arity prep_tycon prep_sort prep sg (t, Ss, x) =
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   283
  test_arity sg (prep_tycon sg t, map (prep_sort sg) Ss, prep sg x);
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   284
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   285
val read_arity = prep_arity Sign.intern_tycon Sign.read_sort Sign.read_sort;
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   286
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
   287
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   288
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
(* 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
   290
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   291
local
423
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
   292
15876
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   293
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
   294
  let
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 sign = Theory.sign_of thy;
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   296
    val (c1, c2) = prep_classrel sign raw_cc;
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   297
    val txt = quote (Sign.string_of_classrel sign [c1, c2]);
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   298
    val _ = message ("Proving class inclusion " ^ txt ^ " ...");
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   299
    val th = Tactic.prove sign [] [] (mk_classrel (c1, c2)) (K tac) handle ERROR =>
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   300
      error ("The error(s) above occurred while trying to prove " ^ txt);
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   301
  in add_classrel_thms [th] thy end;
404
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   302
15876
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   303
fun ext_inst_arity prep_arity raw_arity 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
   304
  let
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   305
    val sign = Theory.sign_of thy;
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   306
    val arity = prep_arity sign raw_arity;
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   307
    val txt = quote (Sign.string_of_arity sign arity);
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   308
    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
   309
    val props = (mk_arities arity);
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   310
    val ths = Tactic.prove_multi sign [] [] props
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   311
        (fn _ => Tactic.smart_conjunction_tac (length props) THEN tac)
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   312
      handle ERROR => error ("The error(s) above occurred while trying to prove " ^ txt);
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   313
  in add_arity_thms ths 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
   314
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   315
in
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   316
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   317
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
   318
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
   319
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
   320
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
   321
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   322
end;
404
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
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
(* 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
   326
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   327
local
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   328
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   329
fun inst_proof mk_prop add_thms inst int theory =
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   330
  theory
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   331
  |> IsarThy.multi_theorem_i Drule.internalK I ProofContext.export_standard
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   332
    ("", [fn (thy, th) => (add_thms [th] thy, th)]) []
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   333
    (map (fn t => (("", []), [(t, ([], []))])) (mk_prop (Theory.sign_of theory) inst)) int;
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   334
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   335
in
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   336
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   337
val instance_subclass_proof =
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   338
  inst_proof (single oo (mk_classrel oo read_classrel)) add_classrel_thms;
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   339
val instance_subclass_proof_i =
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   340
  inst_proof (single oo (mk_classrel oo cert_classrel)) add_classrel_thms;
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   341
val instance_arity_proof = inst_proof (mk_arities oo read_arity) add_arity_thms;
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   342
val instance_arity_proof_i = inst_proof (mk_arities oo cert_arity) add_arity_thms;
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   343
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   344
end;
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   345
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   346
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   347
(* tactics and methods *)
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   348
14605
9de4d64eee3b 'instance' and intro_classes now handle general sorts;
wenzelm
parents: 12876
diff changeset
   349
fun intro_classes_tac facts st =
9de4d64eee3b 'instance' and intro_classes now handle general sorts;
wenzelm
parents: 12876
diff changeset
   350
  (ALLGOALS (Method.insert_tac facts THEN'
9de4d64eee3b 'instance' and intro_classes now handle general sorts;
wenzelm
parents: 12876
diff changeset
   351
      REPEAT_ALL_NEW (resolve_tac (class_axms (Thm.sign_of_thm st))))
9de4d64eee3b 'instance' and intro_classes now handle general sorts;
wenzelm
parents: 12876
diff changeset
   352
    THEN Tactic.distinct_subgoals_tac) st;
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   353
14605
9de4d64eee3b 'instance' and intro_classes now handle general sorts;
wenzelm
parents: 12876
diff changeset
   354
fun default_intro_classes_tac [] = intro_classes_tac []
9de4d64eee3b 'instance' and intro_classes now handle general sorts;
wenzelm
parents: 12876
diff changeset
   355
  | default_intro_classes_tac _ = Tactical.no_tac;    (*no error message!*)
10309
a7f961fb62c6 intro_classes by default;
wenzelm
parents: 10008
diff changeset
   356
a7f961fb62c6 intro_classes by default;
wenzelm
parents: 10008
diff changeset
   357
fun default_tac rules ctxt facts =
14605
9de4d64eee3b 'instance' and intro_classes now handle general sorts;
wenzelm
parents: 12876
diff changeset
   358
  HEADGOAL (Method.some_rule_tac rules ctxt facts) ORELSE
9de4d64eee3b 'instance' and intro_classes now handle general sorts;
wenzelm
parents: 12876
diff changeset
   359
    default_intro_classes_tac facts;
10309
a7f961fb62c6 intro_classes by default;
wenzelm
parents: 10008
diff changeset
   360
15801
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15705
diff changeset
   361
val _ = Context.add_setup [Method.add_methods
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15705
diff changeset
   362
 [("intro_classes", Method.no_args (Method.METHOD intro_classes_tac),
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15705
diff changeset
   363
   "back-chain introduction rules of axiomatic type classes"),
d2f5ca3c048d superceded by Pure.thy and CPure.thy;
wenzelm
parents: 15705
diff changeset
   364
  ("default", Method.thms_ctxt_args (Method.METHOD oo default_tac), "apply some intro/elim rule")]];
10309
a7f961fb62c6 intro_classes by default;
wenzelm
parents: 10008
diff changeset
   365
a7f961fb62c6 intro_classes by default;
wenzelm
parents: 10008
diff changeset
   366
423
a42892e72854 (beta release)
wenzelm
parents: 404
diff changeset
   367
15876
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   368
(** outer syntax **)
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   369
6719
7c2dafc8e801 outer syntax keyword classification;
wenzelm
parents: 6679
diff changeset
   370
local structure P = OuterParse and K = OuterSyntax.Keyword in
3949
c60ff6d0a6b8 fixed goal_XXX;
wenzelm
parents: 3938
diff changeset
   371
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   372
val axclassP =
6719
7c2dafc8e801 outer syntax keyword classification;
wenzelm
parents: 6679
diff changeset
   373
  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
   374
    ((P.name -- Scan.optional ((P.$$$ "\\<subseteq>" || P.$$$ "<") |--
a70df1e5bf10 got rid of explicit marginal comments (now stripped earlier from input);
wenzelm
parents: 12694
diff changeset
   375
        P.!!! (P.list1 P.xname)) []) -- Scan.repeat P.spec_name
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   376
      >> (fn (cls, axs) => Toplevel.theory (#1 o add_axclass cls axs)));
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   377
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   378
val instanceP =
6719
7c2dafc8e801 outer syntax keyword classification;
wenzelm
parents: 6679
diff changeset
   379
  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: 12694
diff changeset
   380
    ((P.xname -- ((P.$$$ "\\<subseteq>" || P.$$$ "<") |-- P.xname) >> instance_subclass_proof ||
14605
9de4d64eee3b 'instance' and intro_classes now handle general sorts;
wenzelm
parents: 12876
diff changeset
   381
      (P.xname -- (P.$$$ "::" |-- P.arity) >> P.triple2) >> instance_arity_proof)
6379
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   382
      >> (Toplevel.print oo Toplevel.theory_to_proof));
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   383
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   384
val _ = OuterSyntax.add_parsers [axclassP, instanceP];
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   385
2b17ff28a6cc theory data;
wenzelm
parents: 6084
diff changeset
   386
end;
3949
c60ff6d0a6b8 fixed goal_XXX;
wenzelm
parents: 3938
diff changeset
   387
15876
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   388
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   389
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   390
(** old-style instantiation proofs **)
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   391
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   392
(* axclass_tac *)
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   393
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   394
(*(1) repeatedly resolve goals of form "OFCLASS(ty, c_class)",
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   395
      try class_trivs first, then "cI" axioms
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   396
  (2) rewrite goals using user supplied definitions
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   397
  (3) repeatedly resolve goals with user supplied non-definitions*)
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   398
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   399
val is_def = Logic.is_equals o #prop o rep_thm;
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   400
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   401
fun axclass_tac thms =
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   402
  let
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   403
    val defs = List.filter is_def thms;
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   404
    val non_defs = filter_out is_def thms;
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   405
  in
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   406
    intro_classes_tac [] THEN
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   407
    TRY (rewrite_goals_tac defs) THEN
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   408
    TRY (REPEAT_FIRST (fn i => assume_tac i ORELSE resolve_tac non_defs i))
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   409
  end;
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   410
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   411
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   412
(* instance declarations *)
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   413
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   414
local
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   415
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   416
fun prove mk_prop str_of sign prop thms usr_tac =
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   417
  (Tactic.prove sign [] [] (mk_prop prop)
15973
wenzelm
parents: 15876
diff changeset
   418
      (K (axclass_tac thms THEN (if_none usr_tac all_tac)))
15876
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   419
    handle ERROR => error ("The error(s) above occurred while trying to prove " ^
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   420
     quote (str_of sign prop))) |> Drule.standard;
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   421
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   422
val prove_subclass = prove mk_classrel (fn sg => fn (c1, c2) =>
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   423
  Pretty.string_of_classrel (Sign.pp sg) [c1, c2]);
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   424
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   425
val prove_arity = prove mk_arity (fn sg => fn (t, Ss, c) =>
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   426
  Pretty.string_of_arity (Sign.pp sg) (t, Ss, [c]));
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   427
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   428
fun witnesses thy names thms =
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   429
  PureThy.get_thmss thy (map (rpair NONE) names) @ thms @ List.filter is_def (map snd (axioms_of thy));
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   430
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   431
in
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   432
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   433
fun add_inst_subclass_x raw_c1_c2 names thms usr_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
   434
  let
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   435
    val sign = Theory.sign_of thy;
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   436
    val (c1, c2) = read_classrel sign raw_c1_c2;
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   437
  in
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   438
    message ("Proving class inclusion " ^ quote (Sign.string_of_classrel sign [c1, c2]) ^ " ...");
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   439
    thy |> add_classrel_thms [prove_subclass sign (c1, c2) (witnesses thy names thms) usr_tac]
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   440
  end;
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   441
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   442
fun add_inst_arity_x raw_arity names thms usr_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
   443
  let
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   444
    val sign = Theory.sign_of thy;
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   445
    val (t, Ss, cs) = read_arity sign raw_arity;
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   446
    val wthms = witnesses thy names thms;
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   447
    fun prove c =
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   448
     (message ("Proving type arity " ^
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   449
        quote (Sign.string_of_arity sign (t, Ss, [c])) ^ " ...");
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   450
        prove_arity sign (t, Ss, c) wthms usr_tac);
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   451
  in add_arity_thms (map prove cs) 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
   452
404
dd3d3d6467db axiomatic type class 'package' for Pure (alpha version);
wenzelm
parents:
diff changeset
   453
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
   454
a67343c6ab2a sane interfaces for tactical instance proofs (do not expand defs of theory, proper handling of sort instances);
wenzelm
parents: 15853
diff changeset
   455
end;