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