src/Pure/Isar/class.ML
author wenzelm
Fri, 19 Oct 2007 20:57:14 +0200
changeset 25104 26b9a7db3386
parent 25103 1ee419a5a30f
child 25146 c2a41f31cacb
permissions -rw-r--r--
tuned interfaces;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
     1
(*  Title:      Pure/Isar/class.ML
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
     2
    ID:         $Id$
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
     3
    Author:     Florian Haftmann, TU Muenchen
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
     4
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
     5
Type classes derived from primitive axclasses and locales.
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
     6
*)
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
     7
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
     8
signature CLASS =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
     9
sig
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    10
  val axclass_cmd: bstring * xstring list
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    11
    -> ((bstring * Attrib.src list) * string list) list
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    12
    -> theory -> class * theory
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    13
  val classrel_cmd: xstring * xstring -> theory -> Proof.state
24914
95cda5dd58d5 added proper subclass concept; improved class target
haftmann
parents: 24901
diff changeset
    14
25002
haftmann
parents: 24981
diff changeset
    15
  val class: bstring -> class list -> Element.context_i Locale.element list
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    16
    -> string list -> theory -> string * Proof.context
25002
haftmann
parents: 24981
diff changeset
    17
  val class_cmd: bstring -> xstring list -> Element.context Locale.element list
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    18
    -> xstring list -> theory -> string * Proof.context
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
    19
  val is_class: theory -> class -> bool
25002
haftmann
parents: 24981
diff changeset
    20
  val these_params: theory -> sort -> (string * (string * typ)) list
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
    21
  val init: class -> Proof.context -> Proof.context
25096
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
    22
  val add_logical_const: string -> Markup.property list
25104
26b9a7db3386 tuned interfaces;
wenzelm
parents: 25103
diff changeset
    23
    -> (string * mixfix) * term -> theory -> theory
25096
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
    24
  val add_syntactic_const: string -> Syntax.mode -> Markup.property list
25104
26b9a7db3386 tuned interfaces;
wenzelm
parents: 25103
diff changeset
    25
    -> (string * mixfix) * term -> theory -> theory
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
    26
  val refresh_syntax: class -> Proof.context -> Proof.context
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    27
  val intro_classes_tac: thm list -> tactic
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    28
  val default_intro_classes_tac: thm list -> tactic
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    29
  val print_classes: theory -> unit
25002
haftmann
parents: 24981
diff changeset
    30
  val uncheck: bool ref
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
    31
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    32
  val instance_arity: (theory -> theory) -> arity list -> theory -> Proof.state
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    33
  val instance: arity list -> ((bstring * Attrib.src list) * term) list
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
    34
    -> (thm list -> theory -> theory)
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    35
    -> theory -> Proof.state
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    36
  val instance_cmd: (bstring * xstring list * xstring) list
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    37
    -> ((bstring * Attrib.src list) * xstring) list
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
    38
    -> (thm list -> theory -> theory)
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    39
    -> theory -> Proof.state
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    40
  val prove_instance: tactic -> arity list
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    41
    -> ((bstring * Attrib.src list) * term) list
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
    42
    -> theory -> thm list * theory
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
    43
  val unoverload: theory -> conv
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
    44
  val overload: theory -> conv
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
    45
  val unoverload_const: theory -> string * typ -> string
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
    46
  val inst_const: theory -> string * string -> string
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
    47
  val param_const: theory -> string -> (string * string) option
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    48
end;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    49
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    50
structure Class : CLASS =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    51
struct
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    52
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    53
(** auxiliary **)
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    54
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
    55
val classN = "class";
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
    56
val introN = "intro";
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
    57
25002
haftmann
parents: 24981
diff changeset
    58
fun prove_interpretation tac prfx_atts expr inst =
haftmann
parents: 24981
diff changeset
    59
  Locale.interpretation_i I prfx_atts expr inst
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    60
  #> Proof.global_terminal_proof
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    61
      (Method.Basic (K (Method.SIMPLE_METHOD tac), Position.none), NONE)
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    62
  #> ProofContext.theory_of;
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    63
25020
f1344290e420 (un)overload: full rewrite;
wenzelm
parents: 25010
diff changeset
    64
fun OF_LAST thm1 thm2 = thm1 RSN (Thm.nprems_of thm2, thm2);
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    65
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    66
fun strip_all_ofclass thy sort =
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    67
  let
24847
bc15dcaed517 replaced AxClass.param_tyvarname by Name.aT;
wenzelm
parents: 24836
diff changeset
    68
    val typ = TVar ((Name.aT, 0), sort);
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    69
    fun prem_inclass t =
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    70
      case Logic.strip_imp_prems t
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    71
       of ofcls :: _ => try Logic.dest_inclass ofcls
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    72
        | [] => NONE;
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    73
    fun strip_ofclass class thm =
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    74
      thm OF (fst o AxClass.of_sort thy (typ, [class])) AxClass.cache;
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    75
    fun strip thm = case (prem_inclass o Thm.prop_of) thm
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    76
     of SOME (_, class) => thm |> strip_ofclass class |> strip
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    77
      | NONE => thm;
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    78
  in strip end;
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    79
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
    80
fun get_remove_global_constraint c thy =
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
    81
  let
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
    82
    val ty = Sign.the_const_constraint thy c;
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
    83
  in
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
    84
    thy
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
    85
    |> Sign.add_const_constraint (c, NONE)
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
    86
    |> pair (c, Logic.unvarifyT ty)
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
    87
  end;
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
    88
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    89
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    90
(** axclass command **)
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    91
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    92
fun axclass_cmd (class, raw_superclasses) raw_specs thy =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    93
  let
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    94
    val ctxt = ProofContext.init thy;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    95
    val superclasses = map (Sign.read_class thy) raw_superclasses;
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    96
    val name_atts = map ((apsnd o map) (Attrib.attribute thy) o fst)
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    97
      raw_specs;
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    98
    val axiomss = ProofContext.read_propp (ctxt, map (map (rpair []) o snd)
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    99
          raw_specs)
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   100
      |> snd
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   101
      |> (map o map) fst;
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   102
  in
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   103
    AxClass.define_class (class, superclasses) []
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   104
      (name_atts ~~ axiomss) thy
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   105
  end;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   106
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   107
local
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   108
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   109
fun gen_instance mk_prop add_thm after_qed insts thy =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   110
  let
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   111
    fun after_qed' results =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   112
      ProofContext.theory ((fold o fold) add_thm results #> after_qed);
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   113
  in
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   114
    thy
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   115
    |> ProofContext.init
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   116
    |> Proof.theorem_i NONE after_qed' ((map (fn t => [(t, [])])
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   117
        o maps (mk_prop thy)) insts)
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   118
  end;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   119
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   120
in
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   121
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   122
val instance_arity =
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   123
  gen_instance (Logic.mk_arities oo Sign.cert_arity) AxClass.add_arity;
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   124
val classrel =
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   125
  gen_instance (single oo (Logic.mk_classrel oo AxClass.cert_classrel))
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   126
    AxClass.add_classrel I o single;
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   127
val classrel_cmd =
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   128
  gen_instance (single oo (Logic.mk_classrel oo AxClass.read_classrel))
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   129
    AxClass.add_classrel I o single;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   130
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   131
end; (*local*)
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   132
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   133
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   134
(** explicit constants for overloaded definitions **)
24304
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   135
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   136
structure InstData = TheoryDataFun
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   137
(
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
   138
  type T = (string * thm) Symtab.table Symtab.table * (string * string) Symtab.table;
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
   139
    (*constant name ~> type constructor ~> (constant name, equation),
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
   140
        constant name ~> (constant name, type constructor)*)
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
   141
  val empty = (Symtab.empty, Symtab.empty);
24304
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   142
  val copy = I;
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   143
  val extend = I;
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
   144
  fun merge _ ((taba1, tabb1), (taba2, tabb2)) =
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
   145
    (Symtab.join (K (Symtab.merge (K true))) (taba1, taba2),
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
   146
      Symtab.merge (K true) (tabb1, tabb2));
24304
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   147
);
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   148
25020
f1344290e420 (un)overload: full rewrite;
wenzelm
parents: 25010
diff changeset
   149
fun inst_thms thy = (Symtab.fold (Symtab.fold (cons o snd o snd) o snd) o fst)
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   150
    (InstData.get thy) [];
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   151
fun add_inst (c, tyco) inst = (InstData.map o apfst
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   152
      o Symtab.map_default (c, Symtab.empty)) (Symtab.update_new (tyco, inst))
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
   153
  #> (InstData.map o apsnd) (Symtab.update_new (fst inst, (c, tyco)));
24304
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   154
25020
f1344290e420 (un)overload: full rewrite;
wenzelm
parents: 25010
diff changeset
   155
fun unoverload thy = MetaSimplifier.rewrite true (inst_thms thy);
f1344290e420 (un)overload: full rewrite;
wenzelm
parents: 25010
diff changeset
   156
fun overload thy = MetaSimplifier.rewrite true (map Thm.symmetric (inst_thms thy));
24304
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   157
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   158
fun inst_const thy (c, tyco) =
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
   159
  (fst o the o Symtab.lookup ((the o Symtab.lookup (fst (InstData.get thy))) c)) tyco;
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
   160
fun unoverload_const thy (c_ty as (c, _)) =
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
   161
  case AxClass.class_of_param thy c
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
   162
   of SOME class => (case Sign.const_typargs thy c_ty
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   163
       of [Type (tyco, _)] => (case Symtab.lookup
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   164
           ((the o Symtab.lookup (fst (InstData.get thy))) c) tyco
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
   165
             of SOME (c, _) => c
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
   166
              | NONE => c)
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
   167
        | [_] => c)
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
   168
    | NONE => c;
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
   169
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
   170
val param_const = Symtab.lookup o snd o InstData.get;
24304
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   171
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   172
fun add_inst_def (class, tyco) (c, ty) thy =
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   173
  let
25024
0615bb9955dd added is_class;
wenzelm
parents: 25020
diff changeset
   174
    val tyco_base = Sign.base_name tyco;
0615bb9955dd added is_class;
wenzelm
parents: 25020
diff changeset
   175
    val name_inst = Sign.base_name class ^ "_" ^ tyco_base ^ "_inst";
0615bb9955dd added is_class;
wenzelm
parents: 25020
diff changeset
   176
    val c_inst_base = Sign.base_name c ^ "_" ^ tyco_base;
24304
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   177
  in
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   178
    thy
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   179
    |> Sign.sticky_prefix name_inst
25020
f1344290e420 (un)overload: full rewrite;
wenzelm
parents: 25010
diff changeset
   180
    |> Sign.declare_const [] (c_inst_base, ty, NoSyn)
24968
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   181
    |-> (fn const as Const (c_inst, _) =>
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   182
      PureThy.add_defs_i false
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   183
        [((Thm.def_name c_inst_base, Logic.mk_equals (const, Const (c, ty))), [])]
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   184
      #-> (fn [thm] => add_inst (c, tyco) (c_inst, Thm.symmetric thm)))
24435
cabf8cd38258 introduces params_of_sort
haftmann
parents: 24423
diff changeset
   185
    |> Sign.restore_naming thy
24304
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   186
  end;
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   187
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   188
fun add_inst_def' (class, tyco) (c, ty) thy =
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
   189
  if case Symtab.lookup (fst (InstData.get thy)) c
24304
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   190
   of NONE => true
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   191
    | SOME tab => is_none (Symtab.lookup tab tyco)
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   192
  then add_inst_def (class, tyco) (c, Logic.unvarifyT ty) thy
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   193
  else thy;
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   194
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   195
fun add_def ((class, tyco), ((name, prop), atts)) thy =
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   196
  let
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   197
    val ((lhs as Const (c, ty), args), rhs) =
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   198
      (apfst Term.strip_comb o Logic.dest_equals) prop;
24701
f8bfd592a6dc no cleverness for instance parameters
haftmann
parents: 24657
diff changeset
   199
    fun (*add_inst' def ([], (Const (c_inst, ty))) =
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
   200
          if forall (fn TFree _ => true | _ => false) (Sign.const_typargs thy (c_inst, ty))
24304
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   201
          then add_inst (c, tyco) (c_inst, def)
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   202
          else add_inst_def (class, tyco) (c, ty)
24701
f8bfd592a6dc no cleverness for instance parameters
haftmann
parents: 24657
diff changeset
   203
      |*) add_inst' _ t = add_inst_def (class, tyco) (c, ty);
24304
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   204
  in
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   205
    thy
25020
f1344290e420 (un)overload: full rewrite;
wenzelm
parents: 25010
diff changeset
   206
    |> PureThy.add_defs_i true [((name, prop), map (Attrib.attribute_i thy) atts)]
24304
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   207
    |-> (fn [def] => add_inst' def (args, rhs) #> pair def)
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   208
  end;
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   209
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   210
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   211
(** instances with implicit parameter handling **)
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   212
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   213
local
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   214
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24701
diff changeset
   215
fun gen_read_def thy prep_att parse_prop ((raw_name, raw_atts), raw_t) =
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   216
  let
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24701
diff changeset
   217
    val ctxt = ProofContext.init thy;
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24701
diff changeset
   218
    val t = parse_prop ctxt raw_t |> Syntax.check_prop ctxt;
24981
4ec3f95190bf dest/cert_def: replaced Pretty.pp by explicit Proof.context;
wenzelm
parents: 24968
diff changeset
   219
    val ((c, ty), _) = Sign.cert_def ctxt t;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   220
    val atts = map (prep_att thy) raw_atts;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   221
    val insts = Consts.typargs (Sign.consts_of thy) (c, ty);
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   222
    val name = case raw_name
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   223
     of "" => NONE
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   224
      | _ => SOME raw_name;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   225
  in (c, (insts, ((name, t), atts))) end;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   226
24707
dfeb98f84e93 Syntax.parse/check/read;
wenzelm
parents: 24701
diff changeset
   227
fun read_def_cmd thy = gen_read_def thy Attrib.intern_src Syntax.parse_prop;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   228
fun read_def thy = gen_read_def thy (K I) (K I);
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   229
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   230
fun gen_instance prep_arity read_def do_proof raw_arities raw_defs after_qed theory =
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   231
  let
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   232
    val arities = map (prep_arity theory) raw_arities;
25020
f1344290e420 (un)overload: full rewrite;
wenzelm
parents: 25010
diff changeset
   233
    val _ = if null arities then error "At least one arity must be given" else ();
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   234
    val _ = case (duplicates (op =) o map #1) arities
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   235
     of [] => ()
25020
f1344290e420 (un)overload: full rewrite;
wenzelm
parents: 25010
diff changeset
   236
      | dupl_tycos => error ("Type constructors occur more than once in arities: "
f1344290e420 (un)overload: full rewrite;
wenzelm
parents: 25010
diff changeset
   237
          ^ commas_quote dupl_tycos);
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   238
    fun get_consts_class tyco ty class =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   239
      let
24968
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   240
        val cs = (these o try (#params o AxClass.get_info theory)) class;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   241
        val subst_ty = map_type_tfree (K ty);
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   242
      in
24304
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   243
        map (fn (c, ty) => (c, ((class, tyco), subst_ty ty))) cs
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   244
      end;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   245
    fun get_consts_sort (tyco, asorts, sort) =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   246
      let
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   247
        val ty = Type (tyco, map (fn (v, sort) => TVar ((v, 0), sort))
24847
bc15dcaed517 replaced AxClass.param_tyvarname by Name.aT;
wenzelm
parents: 24836
diff changeset
   248
          (Name.names Name.context Name.aT asorts))
24731
c25aa6ae64ec Sign.minimize/complete_sort;
wenzelm
parents: 24707
diff changeset
   249
      in maps (get_consts_class tyco ty) (Sign.complete_sort theory sort) end;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   250
    val cs = maps get_consts_sort arities;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   251
    fun mk_typnorm thy (ty, ty_sc) =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   252
      case try (Sign.typ_match thy (Logic.varifyT ty_sc, ty)) Vartab.empty
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   253
       of SOME env => SOME (Logic.varifyT #> Envir.typ_subst_TVars env #> Logic.unvarifyT)
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   254
        | NONE => NONE;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   255
    fun read_defs defs cs thy_read =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   256
      let
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   257
        fun read raw_def cs =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   258
          let
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   259
            val (c, (inst, ((name_opt, t), atts))) = read_def thy_read raw_def;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   260
            val ty = Consts.instance (Sign.consts_of thy_read) (c, inst);
24304
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   261
            val ((class, tyco), ty') = case AList.lookup (op =) cs c
25020
f1344290e420 (un)overload: full rewrite;
wenzelm
parents: 25010
diff changeset
   262
             of NONE => error ("Illegal definition for constant " ^ quote c)
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   263
              | SOME class_ty => class_ty;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   264
            val name = case name_opt
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   265
             of NONE => Thm.def_name (Logic.name_arity (tyco, [], c))
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   266
              | SOME name => name;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   267
            val t' = case mk_typnorm thy_read (ty', ty)
25020
f1344290e420 (un)overload: full rewrite;
wenzelm
parents: 25010
diff changeset
   268
             of NONE => error ("Illegal definition for constant " ^
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   269
                  quote (c ^ "::" ^ setmp show_sorts true
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   270
                    (Sign.string_of_typ thy_read) ty))
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   271
              | SOME norm => map_types norm t
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   272
          in (((class, tyco), ((name, t'), atts)), AList.delete (op =) c cs) end;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   273
      in fold_map read defs cs end;
24304
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   274
    val (defs, other_cs) = read_defs raw_defs cs
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   275
      (fold Sign.primitive_arity arities (Theory.copy theory));
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
   276
    fun after_qed' cs defs =
24766
d0de4e48b526 Sign.add_const_constraint;
wenzelm
parents: 24748
diff changeset
   277
      fold Sign.add_const_constraint (map (apsnd SOME) cs)
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
   278
      #> after_qed defs;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   279
  in
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   280
    theory
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   281
    |> fold_map get_remove_global_constraint (map fst cs |> distinct (op =))
24304
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   282
    ||>> fold_map add_def defs
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   283
    ||> fold (fn (c, ((class, tyco), ty)) => add_inst_def' (class, tyco) (c, ty)) other_cs
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
   284
    |-> (fn (cs, defs) => do_proof (after_qed' cs defs) arities defs)
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   285
  end;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   286
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
   287
fun tactic_proof tac f arities defs =
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   288
  fold (fn arity => AxClass.prove_arity arity tac) arities
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
   289
  #> f
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
   290
  #> pair defs;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   291
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   292
in
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   293
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   294
val instance = gen_instance Sign.cert_arity read_def
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   295
  (fn f => fn arities => fn defs => instance_arity f arities);
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   296
val instance_cmd = gen_instance Sign.read_arity read_def_cmd
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   297
  (fn f => fn arities => fn defs => instance_arity f arities);
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   298
fun prove_instance tac arities defs =
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   299
  gen_instance Sign.cert_arity read_def
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   300
    (tactic_proof tac) arities defs (K I);
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   301
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   302
end; (*local*)
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   303
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   304
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   305
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   306
(** class data **)
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   307
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   308
datatype class_data = ClassData of {
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   309
  consts: (string * string) list
24836
dab06e93ec28 intermediate cleanup
haftmann
parents: 24770
diff changeset
   310
    (*locale parameter ~> constant name*),
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   311
  base_sort: sort,
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   312
  inst: term option list
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   313
    (*canonical interpretation*),
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   314
  morphism: morphism,
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   315
    (*partial morphism of canonical interpretation*)
24657
185502d54c3d fixed wrong syntax treatment in class target
haftmann
parents: 24589
diff changeset
   316
  intro: thm,
185502d54c3d fixed wrong syntax treatment in class target
haftmann
parents: 24589
diff changeset
   317
  defs: thm list,
25096
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   318
  operations: (string * (((term * typ) * (typ * int)) * term)) list
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   319
    (*constant name ~> ((locale term & typ,
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   320
        (constant constraint, instantiaton index of class typ)), locale term for uncheck)*)
24657
185502d54c3d fixed wrong syntax treatment in class target
haftmann
parents: 24589
diff changeset
   321
};
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   322
24657
185502d54c3d fixed wrong syntax treatment in class target
haftmann
parents: 24589
diff changeset
   323
fun rep_class_data (ClassData d) = d;
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   324
fun mk_class_data ((consts, base_sort, inst, morphism, intro),
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   325
    (defs, operations)) =
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   326
  ClassData { consts = consts, base_sort = base_sort, inst = inst,
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   327
    morphism = morphism, intro = intro, defs = defs,
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   328
    operations = operations };
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   329
fun map_class_data f (ClassData { consts, base_sort, inst, morphism, intro,
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   330
    defs, operations }) =
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   331
  mk_class_data (f ((consts, base_sort, inst, morphism, intro),
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   332
    (defs, operations)));
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   333
fun merge_class_data _ (ClassData { consts = consts,
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   334
    base_sort = base_sort, inst = inst, morphism = morphism, intro = intro,
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   335
    defs = defs1, operations = operations1 },
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   336
  ClassData { consts = _, base_sort = _, inst = _, morphism = _, intro = _,
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   337
    defs = defs2, operations = operations2 }) =
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   338
  mk_class_data ((consts, base_sort, inst, morphism, intro),
24914
95cda5dd58d5 added proper subclass concept; improved class target
haftmann
parents: 24901
diff changeset
   339
    (Thm.merge_thms (defs1, defs2),
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   340
      AList.merge (op =) (K true) (operations1, operations2)));
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   341
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   342
structure ClassData = TheoryDataFun
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   343
(
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   344
  type T = class_data Graph.T
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   345
  val empty = Graph.empty;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   346
  val copy = I;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   347
  val extend = I;
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   348
  fun merge _ = Graph.join merge_class_data;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   349
);
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   350
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   351
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   352
(* queries *)
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   353
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   354
val lookup_class_data = Option.map rep_class_data oo try o Graph.get_node o ClassData.get;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   355
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   356
fun the_class_data thy class = case lookup_class_data thy class
25020
f1344290e420 (un)overload: full rewrite;
wenzelm
parents: 25010
diff changeset
   357
 of NONE => error ("Undeclared class " ^ quote class)
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   358
  | SOME data => data;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   359
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   360
val is_class = is_some oo lookup_class_data;
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   361
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   362
val ancestry = Graph.all_succs o ClassData.get;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   363
25002
haftmann
parents: 24981
diff changeset
   364
fun these_params thy =
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   365
  let
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   366
    fun params class =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   367
      let
24930
cc2e0e8c81af renamed AxClass.get_definition to AxClass.get_info (again);
wenzelm
parents: 24920
diff changeset
   368
        val const_typs = (#params o AxClass.get_info thy) class;
24657
185502d54c3d fixed wrong syntax treatment in class target
haftmann
parents: 24589
diff changeset
   369
        val const_names = (#consts o the_class_data thy) class;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   370
      in
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   371
        (map o apsnd) (fn c => (c, (the o AList.lookup (op =) const_typs) c)) const_names
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   372
      end;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   373
  in maps params o ancestry thy end;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   374
24657
185502d54c3d fixed wrong syntax treatment in class target
haftmann
parents: 24589
diff changeset
   375
fun these_defs thy = maps (these o Option.map #defs o lookup_class_data thy) o ancestry thy;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   376
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   377
fun morphism thy = #morphism o the_class_data thy;
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   378
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   379
fun these_intros thy =
24657
185502d54c3d fixed wrong syntax treatment in class target
haftmann
parents: 24589
diff changeset
   380
  Graph.fold (fn (_, (data, _)) => insert Thm.eq_thm ((#intro o rep_class_data) data))
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   381
    (ClassData.get thy) [];
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   382
24836
dab06e93ec28 intermediate cleanup
haftmann
parents: 24770
diff changeset
   383
fun these_operations thy =
dab06e93ec28 intermediate cleanup
haftmann
parents: 24770
diff changeset
   384
  maps (#operations o the_class_data thy) o ancestry thy;
24657
185502d54c3d fixed wrong syntax treatment in class target
haftmann
parents: 24589
diff changeset
   385
25002
haftmann
parents: 24981
diff changeset
   386
fun local_operation thy = AList.lookup (op =) o these_operations thy;
haftmann
parents: 24981
diff changeset
   387
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   388
fun sups_base_sort thy sort =
24901
d3cbf79769b9 added first version of user-space type system for class target
haftmann
parents: 24847
diff changeset
   389
  let
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   390
    val sups = filter (is_class thy) sort
24901
d3cbf79769b9 added first version of user-space type system for class target
haftmann
parents: 24847
diff changeset
   391
      |> Sign.minimize_sort thy;
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   392
    val base_sort = case sups
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   393
     of _ :: _ => maps (#base_sort o the_class_data thy) sups
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   394
          |> Sign.minimize_sort thy
24901
d3cbf79769b9 added first version of user-space type system for class target
haftmann
parents: 24847
diff changeset
   395
      | [] => sort;
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   396
  in (sups, base_sort) end;
24901
d3cbf79769b9 added first version of user-space type system for class target
haftmann
parents: 24847
diff changeset
   397
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   398
fun print_classes thy =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   399
  let
24920
2a45e400fdad generic Syntax.pretty/string_of operations;
wenzelm
parents: 24914
diff changeset
   400
    val ctxt = ProofContext.init thy;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   401
    val algebra = Sign.classes_of thy;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   402
    val arities =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   403
      Symtab.empty
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   404
      |> Symtab.fold (fn (tyco, arities) => fold (fn (class, _) =>
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   405
           Symtab.map_default (class, []) (insert (op =) tyco)) arities)
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   406
             ((#arities o Sorts.rep_algebra) algebra);
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   407
    val the_arities = these o Symtab.lookup arities;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   408
    fun mk_arity class tyco =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   409
      let
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   410
        val Ss = Sorts.mg_domain algebra tyco [class];
24920
2a45e400fdad generic Syntax.pretty/string_of operations;
wenzelm
parents: 24914
diff changeset
   411
      in Syntax.pretty_arity ctxt (tyco, Ss, [class]) end;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   412
    fun mk_param (c, ty) = Pretty.str (Sign.extern_const thy c ^ " :: "
24920
2a45e400fdad generic Syntax.pretty/string_of operations;
wenzelm
parents: 24914
diff changeset
   413
      ^ setmp show_sorts false (Syntax.string_of_typ ctxt o Type.strip_sorts) ty);
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   414
    fun mk_entry class = (Pretty.block o Pretty.fbreaks o map_filter I) [
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   415
      (SOME o Pretty.str) ("class " ^ Sign.extern_class thy class ^ ":"),
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   416
      (SOME o Pretty.block) [Pretty.str "supersort: ",
24920
2a45e400fdad generic Syntax.pretty/string_of operations;
wenzelm
parents: 24914
diff changeset
   417
        (Syntax.pretty_sort ctxt o Sign.minimize_sort thy o Sign.super_classes thy) class],
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   418
      if is_class thy class then (SOME o Pretty.str)
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   419
        ("locale: " ^ Locale.extern thy class) else NONE,
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   420
      ((fn [] => NONE | ps => (SOME o Pretty.block o Pretty.fbreaks)
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   421
          (Pretty.str "parameters:" :: ps)) o map mk_param
24930
cc2e0e8c81af renamed AxClass.get_definition to AxClass.get_info (again);
wenzelm
parents: 24920
diff changeset
   422
        o these o Option.map #params o try (AxClass.get_info thy)) class,
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   423
      (SOME o Pretty.block o Pretty.breaks) [
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   424
        Pretty.str "instances:",
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   425
        Pretty.list "" "" (map (mk_arity class) (the_arities class))
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   426
      ]
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   427
    ]
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   428
  in
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   429
    (Pretty.writeln o Pretty.chunks o separate (Pretty.str "")
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   430
      o map mk_entry o Sorts.all_classes) algebra
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   431
  end;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   432
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   433
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   434
(* updaters *)
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   435
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   436
fun add_class_data ((class, superclasses), (cs, base_sort, insttab, phi, intro)) thy =
25002
haftmann
parents: 24981
diff changeset
   437
  let
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   438
    val inst = map
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   439
      (SOME o the o Symtab.lookup insttab o fst o fst)
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   440
        (Locale.parameters_of_expr thy (Locale.Locale class));
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   441
    val operations = map (fn (v_ty as (_, ty'), (c, ty)) =>
25096
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   442
      (c, (((Free v_ty, ty'), (Logic.varifyT ty, 0)), Free v_ty))) cs;
25002
haftmann
parents: 24981
diff changeset
   443
    val cs = (map o pairself) fst cs;
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   444
    val add_class = Graph.new_node (class,
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   445
        mk_class_data ((cs, base_sort, inst, phi, intro), ([], operations)))
25002
haftmann
parents: 24981
diff changeset
   446
      #> fold (curry Graph.add_edge class) superclasses;
haftmann
parents: 24981
diff changeset
   447
  in
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   448
    ClassData.map add_class thy
25002
haftmann
parents: 24981
diff changeset
   449
  end;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   450
25096
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   451
fun register_operation class ((c, (dict, dict_rev)), some_def) thy =
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   452
  let
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   453
    val ty = Sign.the_const_constraint thy c;
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   454
    val typargs = Sign.const_typargs thy (c, ty);
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   455
    val typidx = find_index (fn TVar ((v, _), _) => Name.aT = v | _ => false) typargs;
25096
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   456
    val ty' = Term.fastype_of dict;
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   457
  in
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   458
    thy
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   459
    |> (ClassData.map o Graph.map_node class o map_class_data o apsnd)
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   460
      (fn (defs, operations) =>
25096
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   461
        (fold cons (the_list some_def) defs,
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   462
          (c, (((dict, ty'), (ty, typidx)), dict_rev)) :: operations))
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   463
  end;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   464
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   465
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   466
(** rule calculation, tactics and methods **)
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   467
25024
0615bb9955dd added is_class;
wenzelm
parents: 25020
diff changeset
   468
val class_prefix = Logic.const_of_class o Sign.base_name;
0615bb9955dd added is_class;
wenzelm
parents: 25020
diff changeset
   469
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   470
fun calculate_morphism class cs =
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   471
  let
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   472
    val subst_typ = Term.map_type_tfree (fn var as (v, sort) =>
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   473
      if v = Name.aT then TVar ((v, 0), [class]) else TVar ((v, 0), sort));
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   474
    fun subst_aterm (t as Free (v, ty)) = (case AList.lookup (op =) cs v
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   475
         of SOME (c, _) => Const (c, ty)
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   476
          | NONE => t)
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   477
      | subst_aterm t = t;
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   478
    val subst_term = map_aterms subst_aterm #> map_types subst_typ;
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   479
  in
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   480
    Morphism.identity
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   481
    $> Morphism.term_morphism subst_term
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   482
    $> Morphism.typ_morphism subst_typ
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   483
  end;
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   484
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   485
fun class_intro thy class sups =
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   486
  let
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   487
    fun class_elim class =
25020
f1344290e420 (un)overload: full rewrite;
wenzelm
parents: 25010
diff changeset
   488
      case (#axioms o AxClass.get_info thy) class
f1344290e420 (un)overload: full rewrite;
wenzelm
parents: 25010
diff changeset
   489
       of [thm] => SOME (Drule.unconstrainTs thm)
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   490
        | [] => NONE;
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   491
    val pred_intro = case Locale.intros thy class
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   492
     of ([ax_intro], [intro]) => intro |> OF_LAST ax_intro |> SOME
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   493
      | ([intro], []) => SOME intro
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   494
      | ([], [intro]) => SOME intro
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   495
      | _ => NONE;
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   496
    val pred_intro' = pred_intro
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   497
      |> Option.map (fn intro => intro OF map_filter class_elim sups);
24930
cc2e0e8c81af renamed AxClass.get_definition to AxClass.get_info (again);
wenzelm
parents: 24920
diff changeset
   498
    val class_intro = (#intro o AxClass.get_info thy) class;
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   499
    val raw_intro = case pred_intro'
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   500
     of SOME pred_intro => class_intro |> OF_LAST pred_intro
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   501
      | NONE => class_intro;
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   502
    val sort = Sign.super_classes thy class;
24847
bc15dcaed517 replaced AxClass.param_tyvarname by Name.aT;
wenzelm
parents: 24836
diff changeset
   503
    val typ = TVar ((Name.aT, 0), sort);
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   504
    val defs = these_defs thy sups;
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   505
  in
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   506
    raw_intro
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   507
    |> Drule.instantiate' [SOME (Thm.ctyp_of thy typ)] []
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   508
    |> strip_all_ofclass thy sort
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   509
    |> Thm.strip_shyps
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   510
    |> MetaSimplifier.rewrite_rule defs
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   511
    |> Drule.unconstrainTs
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   512
  end;
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   513
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   514
fun class_interpretation class facts defs thy =
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   515
  let
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   516
    val params = these_params thy [class];
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   517
    val inst = (#inst o the_class_data thy) class;
25020
f1344290e420 (un)overload: full rewrite;
wenzelm
parents: 25010
diff changeset
   518
    val tac = ALLGOALS (ProofContext.fact_tac facts);
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   519
    val prfx = class_prefix class;
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   520
  in
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   521
    thy
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   522
    |> fold_map (get_remove_global_constraint o fst o snd) params
25094
ba43514068fd Interpretation equations may have name and/or attribute.
ballarin
parents: 25083
diff changeset
   523
    ||> prove_interpretation tac ((false, prfx), []) (Locale.Locale class)
ba43514068fd Interpretation equations may have name and/or attribute.
ballarin
parents: 25083
diff changeset
   524
          (inst, map (fn def => (("", []), def)) defs)
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   525
    |-> (fn cs => fold (Sign.add_const_constraint o apsnd SOME) cs)
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   526
  end;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   527
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   528
fun intro_classes_tac facts st =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   529
  let
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   530
    val thy = Thm.theory_of_thm st;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   531
    val classes = Sign.all_classes thy;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   532
    val class_trivs = map (Thm.class_triv thy) classes;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   533
    val class_intros = these_intros thy;
24930
cc2e0e8c81af renamed AxClass.get_definition to AxClass.get_info (again);
wenzelm
parents: 24920
diff changeset
   534
    val axclass_intros = map_filter (try (#intro o AxClass.get_info thy)) classes;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   535
  in
24930
cc2e0e8c81af renamed AxClass.get_definition to AxClass.get_info (again);
wenzelm
parents: 24920
diff changeset
   536
    (ALLGOALS (Method.insert_tac facts THEN'
cc2e0e8c81af renamed AxClass.get_definition to AxClass.get_info (again);
wenzelm
parents: 24920
diff changeset
   537
      REPEAT_ALL_NEW (resolve_tac (class_trivs @ class_intros @ axclass_intros)))
cc2e0e8c81af renamed AxClass.get_definition to AxClass.get_info (again);
wenzelm
parents: 24920
diff changeset
   538
    THEN Tactic.distinct_subgoals_tac) st
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   539
  end;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   540
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   541
fun default_intro_classes_tac [] = intro_classes_tac []
24930
cc2e0e8c81af renamed AxClass.get_definition to AxClass.get_info (again);
wenzelm
parents: 24920
diff changeset
   542
  | default_intro_classes_tac _ = no_tac;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   543
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   544
fun default_tac rules ctxt facts =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   545
  HEADGOAL (Method.some_rule_tac rules ctxt facts) ORELSE
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   546
    default_intro_classes_tac facts;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   547
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   548
val _ = Context.add_setup (Method.add_methods
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   549
 [("intro_classes", Method.no_args (Method.METHOD intro_classes_tac),
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   550
    "back-chain introduction rules of classes"),
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   551
  ("default", Method.thms_ctxt_args (Method.METHOD oo default_tac),
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   552
    "apply some intro/elim rule")]);
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   553
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   554
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   555
(** classes and class target **)
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   556
25002
haftmann
parents: 24981
diff changeset
   557
(* class context syntax *)
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   558
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   559
structure ClassSyntax = ProofDataFun(
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   560
  type T = {
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   561
    constraints: (string * typ) list,
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   562
    base_sort: sort,
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   563
    local_operation: string * typ -> (typ * term) option,
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   564
    rews: (term * term) list,
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   565
    passed: bool
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   566
  } option;
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   567
  fun init _ = NONE;
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   568
);
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   569
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   570
fun synchronize_syntax thy sups base_sort ctxt =
24914
95cda5dd58d5 added proper subclass concept; improved class target
haftmann
parents: 24901
diff changeset
   571
  let
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   572
    val operations = these_operations thy sups;
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   573
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   574
    (* constraints *)
25096
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   575
    fun local_constraint (c, ((_, (ty, _)),_ )) =
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   576
      let
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   577
        val ty' = ty
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   578
          |> map_atyps (fn ty as TVar ((v, 0), _) =>
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   579
               if v = Name.aT then TVar ((v, 0), base_sort) else ty)
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   580
          |> SOME;
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   581
      in (c, ty') end
25096
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   582
    val constraints = (map o apsnd) (fst o snd o fst) operations;
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   583
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   584
    (* check phase *)
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   585
    val typargs = Consts.typargs (ProofContext.consts_of ctxt);
25096
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   586
    fun check_const (c, ty) (((t, _), (_, idx)), _) =
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   587
      ((nth (typargs (c, ty)) idx), t);
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   588
    fun local_operation (c_ty as (c, _)) = AList.lookup (op =) operations c
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   589
      |> Option.map (check_const c_ty);
25002
haftmann
parents: 24981
diff changeset
   590
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   591
    (* uncheck phase *)
25096
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   592
    val proto_rews = map (fn (c, (((t, ty), _), _)) => (t, Const (c, ty))) operations;
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   593
    fun rew_app f (t1 $ t2) = rew_app f t1 $ f t2
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   594
      | rew_app f t = t;
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   595
    val rews = (map o apfst o rew_app)
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   596
      (Pattern.rewrite_term thy proto_rews []) proto_rews;
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   597
  in
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   598
    ctxt
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   599
    |> fold (ProofContext.add_const_constraint o local_constraint) operations
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   600
    |> ClassSyntax.map (K (SOME {
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   601
        constraints = constraints,
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   602
        base_sort = base_sort,
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   603
        local_operation = local_operation,
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   604
        rews = rews,
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   605
        passed = false
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   606
      }))
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   607
  end;
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   608
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   609
fun refresh_syntax class ctxt =
25002
haftmann
parents: 24981
diff changeset
   610
  let
haftmann
parents: 24981
diff changeset
   611
    val thy = ProofContext.theory_of ctxt;
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   612
    val base_sort = (#base_sort o the_class_data thy) class;
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   613
  in synchronize_syntax thy [class] base_sort ctxt end;
24914
95cda5dd58d5 added proper subclass concept; improved class target
haftmann
parents: 24901
diff changeset
   614
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   615
val mark_passed = (ClassSyntax.map o Option.map)
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   616
  (fn { constraints, base_sort, local_operation, rews, passed } =>
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   617
    { constraints = constraints, base_sort = base_sort,
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   618
      local_operation = local_operation, rews = rews, passed = true });
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   619
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   620
fun sort_term_check ts ctxt =
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   621
  let
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   622
    val { constraints, base_sort, local_operation, passed, ... } =
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   623
      the (ClassSyntax.get ctxt);
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   624
    fun check_typ (c, ty) (TFree (v, _), t) = if v = Name.aT
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   625
          then apfst (AList.update (op =) ((c, ty), t)) else I
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   626
      | check_typ (c, ty) (TVar (vi, _), t) = if TypeInfer.is_param vi
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   627
          then apfst (AList.update (op =) ((c, ty), t))
24901
d3cbf79769b9 added first version of user-space type system for class target
haftmann
parents: 24847
diff changeset
   628
            #> apsnd (insert (op =) vi) else I
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   629
      | check_typ _ _ = I;
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   630
    fun add_const (Const c_ty) = Option.map (check_typ c_ty) (local_operation c_ty)
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   631
          |> the_default I
24901
d3cbf79769b9 added first version of user-space type system for class target
haftmann
parents: 24847
diff changeset
   632
      | add_const _ = I;
d3cbf79769b9 added first version of user-space type system for class target
haftmann
parents: 24847
diff changeset
   633
    val (cs, typarams) = (fold o fold_aterms) add_const ts ([], []);
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   634
    val subst_typ = map_type_tvar (fn var as (vi, _) =>
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   635
      if member (op =) typarams vi then TFree (Name.aT, base_sort) else TVar var);
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   636
    val subst_term = map_aterms
24901
d3cbf79769b9 added first version of user-space type system for class target
haftmann
parents: 24847
diff changeset
   637
        (fn t as Const (c, ty) => the_default t (AList.lookup (op =) cs (c, ty)) | t => t)
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   638
          #> map_types subst_typ;
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   639
    val ts' = map subst_term ts;
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   640
  in if eq_list (op aconv) (ts, ts') andalso passed then NONE
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   641
  else
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   642
    ctxt
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   643
    |> fold (ProofContext.add_const_constraint o apsnd SOME) constraints
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   644
    |> mark_passed
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   645
    |> pair ts'
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   646
    |> SOME
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   647
  end;
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   648
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   649
val uncheck = ref true;
25002
haftmann
parents: 24981
diff changeset
   650
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   651
fun sort_term_uncheck ts ctxt =
25002
haftmann
parents: 24981
diff changeset
   652
  let
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   653
    (*FIXME abbreviations*)
25002
haftmann
parents: 24981
diff changeset
   654
    val thy = ProofContext.theory_of ctxt;
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   655
    val rews = (#rews o the o ClassSyntax.get) ctxt;
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   656
    val ts' = if ! uncheck
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   657
      then map (Pattern.rewrite_term thy rews []) ts else ts;
25060
17c313217998 Syntax.(un)check: explicit result option;
wenzelm
parents: 25038
diff changeset
   658
  in if eq_list (op aconv) (ts, ts') then NONE else SOME (ts', ctxt) end;
25002
haftmann
parents: 24981
diff changeset
   659
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   660
fun init_ctxt thy sups base_sort ctxt =
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   661
  ctxt
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   662
  |> Variable.declare_term
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   663
      (Logic.mk_type (TFree (Name.aT, base_sort)))
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   664
  |> synchronize_syntax thy sups base_sort
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   665
  |> Context.proof_map (
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   666
      Syntax.add_term_check 0 "class" sort_term_check
25103
haftmann
parents: 25096
diff changeset
   667
      #> Syntax.add_term_uncheck 0 "class" sort_term_uncheck)
24901
d3cbf79769b9 added first version of user-space type system for class target
haftmann
parents: 24847
diff changeset
   668
d3cbf79769b9 added first version of user-space type system for class target
haftmann
parents: 24847
diff changeset
   669
fun init class ctxt =
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   670
  let
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   671
    val thy = ProofContext.theory_of ctxt;
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   672
  in
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   673
    init_ctxt thy [class] ((#base_sort o the_class_data thy) class) ctxt
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   674
  end;
24914
95cda5dd58d5 added proper subclass concept; improved class target
haftmann
parents: 24901
diff changeset
   675
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   676
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   677
(* class definition *)
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   678
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   679
local
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   680
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   681
fun gen_class_spec prep_class prep_expr process_expr thy raw_supclasses raw_includes_elems =
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   682
  let
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   683
    val supclasses = map (prep_class thy) raw_supclasses;
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   684
    val (sups, base_sort) = sups_base_sort thy supclasses;
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   685
    val supsort = Sign.minimize_sort thy supclasses;
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   686
    val suplocales = map Locale.Locale sups;
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   687
    val (raw_elems, includes) = fold_rev (fn Locale.Elem e => apfst (cons e)
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   688
      | Locale.Expr i => apsnd (cons (prep_expr thy i))) raw_includes_elems ([], []);
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   689
    val supexpr = Locale.Merge suplocales;
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   690
    val supparams = (map fst o Locale.parameters_of_expr thy) supexpr;
25002
haftmann
parents: 24981
diff changeset
   691
    val supconsts = AList.make (the o AList.lookup (op =) (these_params thy sups))
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   692
      (map fst supparams);
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   693
    val mergeexpr = Locale.Merge (suplocales @ includes);
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   694
    val constrain = Element.Constrains ((map o apsnd o map_atyps)
24847
bc15dcaed517 replaced AxClass.param_tyvarname by Name.aT;
wenzelm
parents: 24836
diff changeset
   695
      (fn TFree (_, sort) => TFree (Name.aT, sort)) supparams);
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   696
  in
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   697
    ProofContext.init thy
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   698
    |> Locale.cert_expr supexpr [constrain]
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   699
    |> snd
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   700
    |> init_ctxt thy sups base_sort
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   701
    |> process_expr Locale.empty raw_elems
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   702
    |> fst
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   703
    |> (fn elems => ((((sups, supconsts), (supsort, base_sort, mergeexpr)),
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   704
          (*FIXME*) if null includes then constrain :: elems else elems)))
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   705
  end;
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   706
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   707
val read_class_spec = gen_class_spec Sign.intern_class Locale.intern_expr Locale.read_expr;
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   708
val check_class_spec = gen_class_spec (K I) (K I) Locale.cert_expr;
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   709
24968
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   710
fun define_class_params (name, raw_superclasses) raw_consts raw_dep_axioms other_consts thy =
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   711
  let
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   712
    val superclasses = map (Sign.certify_class thy) raw_superclasses;
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   713
    val consts = (map o apfst o apsnd) (Sign.certify_typ thy) raw_consts;
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   714
    fun add_const ((c, ty), syn) =
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   715
      Sign.declare_const [] (c, Type.strip_sorts ty, syn) #>> Term.dest_Const;
24968
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   716
    fun mk_axioms cs thy =
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   717
      raw_dep_axioms thy cs
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   718
      |> (map o apsnd o map) (Sign.cert_prop thy)
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   719
      |> rpair thy;
25002
haftmann
parents: 24981
diff changeset
   720
    fun constrain_typs class = (map o apsnd o Term.map_type_tfree)
haftmann
parents: 24981
diff changeset
   721
      (fn (v, _) => TFree (v, [class]))
24968
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   722
  in
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   723
    thy
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   724
    |> Sign.add_path (Logic.const_of_class name)
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   725
    |> fold_map add_const consts
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   726
    ||> Sign.restore_naming thy
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   727
    |-> (fn cs => mk_axioms cs
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   728
    #-> (fn axioms_prop => AxClass.define_class (name, superclasses)
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   729
           (map fst cs @ other_consts) axioms_prop
25002
haftmann
parents: 24981
diff changeset
   730
    #-> (fn class => `(fn _ => constrain_typs class cs)
haftmann
parents: 24981
diff changeset
   731
    #-> (fn cs' => `(fn thy => AxClass.get_info thy class)
haftmann
parents: 24981
diff changeset
   732
    #-> (fn {axioms, ...} => fold (Sign.add_const_constraint o apsnd SOME) cs'
haftmann
parents: 24981
diff changeset
   733
    #> pair (class, (cs', axioms)))))))
24968
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   734
  end;
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   735
25002
haftmann
parents: 24981
diff changeset
   736
fun gen_class prep_spec prep_param bname
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   737
    raw_supclasses raw_includes_elems raw_other_consts thy =
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   738
  let
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   739
    val class = Sign.full_name thy bname;
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   740
    val (((sups, supconsts), (supsort, base_sort, mergeexpr)), elems_syn) =
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   741
      prep_spec thy raw_supclasses raw_includes_elems;
24968
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   742
    val other_consts = map (tap (Sign.the_const_type thy) o prep_param thy) raw_other_consts;
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   743
    fun mk_inst class param_names cs =
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   744
      Symtab.empty
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   745
      |> fold2 (fn v => fn (c, ty) => Symtab.update (v, Const
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   746
           (c, Term.map_type_tfree (fn (v, _) => TFree (v, [class])) ty))) param_names cs;
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   747
    fun fork_syntax (Element.Fixes xs) =
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   748
          fold_map (fn (c, ty, syn) => cons (c, syn) #> pair (c, ty, NoSyn)) xs
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   749
          #>> Element.Fixes
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   750
      | fork_syntax x = pair x;
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   751
    val (elems, global_syn) = fold_map fork_syntax elems_syn [];
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   752
    fun globalize (c, ty) = 
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   753
      ((c, Term.map_type_tfree (K (TFree (Name.aT, base_sort))) ty),
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   754
        (the_default NoSyn o AList.lookup (op =) global_syn) c);
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   755
    fun extract_params thy =
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   756
      let
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   757
        val params = map fst (Locale.parameters_of thy class);
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   758
      in
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   759
        (params, (map globalize o snd o chop (length supconsts)) params)
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   760
      end;
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   761
    fun extract_assumes params thy cs =
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   762
      let
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   763
        val consts = supconsts @ (map (fst o fst) params ~~ cs);
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   764
        fun subst (Free (c, ty)) =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   765
              Const ((fst o the o AList.lookup (op =) consts) c, ty)
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   766
          | subst t = t;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   767
        fun prep_asm ((name, atts), ts) =
25024
0615bb9955dd added is_class;
wenzelm
parents: 25020
diff changeset
   768
          ((Sign.base_name name, map (Attrib.attribute_i thy) atts),
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   769
            (map o map_aterms) subst ts);
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   770
      in
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   771
        Locale.global_asms_of thy class
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   772
        |> map prep_asm
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   773
      end;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   774
  in
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   775
    thy
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   776
    |> Locale.add_locale_i (SOME "") bname mergeexpr elems
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   777
    |> snd
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   778
    |> ProofContext.theory (`extract_params
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   779
      #-> (fn (all_params, params) =>
24968
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   780
        define_class_params (bname, supsort) params
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   781
          (extract_assumes params) other_consts
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   782
      #-> (fn (_, (consts, axioms)) =>
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   783
        `(fn thy => class_intro thy class sups)
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   784
      #-> (fn class_intro =>
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   785
        PureThy.note_thmss_qualified "" (NameSpace.append class classN)
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   786
          [((introN, []), [([class_intro], [])])]
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   787
      #-> (fn [(_, [class_intro])] =>
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   788
        add_class_data ((class, sups),
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   789
          (map fst params ~~ consts, base_sort,
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   790
            mk_inst class (map fst all_params) (map snd supconsts @ consts),
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   791
              calculate_morphism class (supconsts @ (map (fst o fst) params ~~ consts)), class_intro))
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   792
      #> class_interpretation class axioms []
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   793
      )))))
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   794
    |> pair class
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   795
  end;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   796
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   797
in
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   798
24968
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   799
val class_cmd = gen_class read_class_spec ((#1 o Term.dest_Const) oo Sign.read_const);
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   800
val class = gen_class check_class_spec (K I);
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   801
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   802
end; (*local*)
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   803
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   804
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   805
(* definition in class target *)
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   806
25096
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   807
fun add_logical_const class pos ((c, mx), dict) thy =
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   808
  let
25024
0615bb9955dd added is_class;
wenzelm
parents: 25020
diff changeset
   809
    val prfx = class_prefix class;
0615bb9955dd added is_class;
wenzelm
parents: 25020
diff changeset
   810
    val thy' = thy |> Sign.add_path prfx;
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   811
    val phi = morphism thy' class;
25024
0615bb9955dd added is_class;
wenzelm
parents: 25020
diff changeset
   812
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   813
    val c' = Sign.full_name thy' c;
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   814
    val dict' = (map_types Logic.unvarifyT o Morphism.term phi) dict;
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   815
    val ty' = Term.fastype_of dict';
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   816
    val ty'' = Type.strip_sorts ty';
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   817
    val def_eq = Logic.mk_equals (Const (c', ty'), dict');
25096
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   818
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   819
    val c'' = NameSpace.full (Sign.naming_of thy' |> NameSpace.add_path prfx) c;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   820
  in
25024
0615bb9955dd added is_class;
wenzelm
parents: 25020
diff changeset
   821
    thy'
25096
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   822
    |> Sign.hide_consts_i false [c''] (*FIXME*)
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   823
    |> Sign.declare_const pos (c, ty'', mx) |> snd
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   824
    |> Thm.add_def false (c, def_eq)
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   825
    |>> Thm.symmetric
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   826
    |-> (fn def => class_interpretation class [def] [Thm.prop_of def]
25096
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   827
          #> register_operation class ((c', (dict, dict')), SOME (Thm.varifyT def)))
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   828
    |> Sign.restore_naming thy
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   829
    |> Sign.add_const_constraint (c', SOME ty')
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   830
  end;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   831
24901
d3cbf79769b9 added first version of user-space type system for class target
haftmann
parents: 24847
diff changeset
   832
d3cbf79769b9 added first version of user-space type system for class target
haftmann
parents: 24847
diff changeset
   833
(* abbreviation in class target *)
d3cbf79769b9 added first version of user-space type system for class target
haftmann
parents: 24847
diff changeset
   834
25096
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   835
fun expand_aux_abbrev thy class t =
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   836
  let
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   837
    val rews = map (Logic.dest_equals o Thm.prop_of) (these_defs thy [class]);
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   838
    val (head, ts) = strip_comb t;
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   839
    val tys = (fst o chop (length ts) o fst o strip_type o Term.fastype_of) head;
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   840
    val head' = head
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   841
      |> Pattern.eta_long tys
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   842
      |> Consts.certify (Sign.pp thy) (Sign.tsig_of thy) true (Sign.consts_of thy);
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   843
    val ts' = ts
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   844
      |> map (Pattern.rewrite_term thy rews []);
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   845
  in Term.betapplys (head', ts') end;
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   846
25103
haftmann
parents: 25096
diff changeset
   847
fun fold_aux_defs thy class =
haftmann
parents: 25096
diff changeset
   848
  let
haftmann
parents: 25096
diff changeset
   849
    val rews = map (Logic.dest_equals o Thm.prop_of) (these_defs thy [class])
haftmann
parents: 25096
diff changeset
   850
  in Pattern.rewrite_term thy rews [] end;
haftmann
parents: 25096
diff changeset
   851
haftmann
parents: 25096
diff changeset
   852
fun add_syntactic_const class prmode pos ((c, mx), rhs) thy =
24836
dab06e93ec28 intermediate cleanup
haftmann
parents: 24770
diff changeset
   853
  let
25024
0615bb9955dd added is_class;
wenzelm
parents: 25020
diff changeset
   854
    val prfx = class_prefix class;
25096
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   855
    val thy' = thy |> Sign.add_path prfx;
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   856
    val phi = morphism thy class;
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   857
25096
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   858
    val c' = Sign.full_name thy' c;
25103
haftmann
parents: 25096
diff changeset
   859
    val dict' = (expand_aux_abbrev thy' class o Morphism.term phi) rhs;
25096
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   860
    val dict'' = map_types Logic.unvarifyT dict';
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   861
    val ty' = Term.fastype_of dict'';
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   862
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   863
    val c'' = NameSpace.full (Sign.naming_of thy' |> NameSpace.add_path prfx) c;
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   864
    val ty'' = (Type.strip_sorts o Logic.unvarifyT) (Sign.the_const_constraint thy' c'');
24836
dab06e93ec28 intermediate cleanup
haftmann
parents: 24770
diff changeset
   865
  in
25096
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   866
    thy'
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   867
    |> Sign.add_const_constraint (c'', SOME ty'') (*FIXME*)
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   868
    |> Sign.hide_consts_i false [c''] (*FIXME*)
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   869
    |> Sign.add_abbrev (#1 prmode) pos (c, map_types Type.strip_sorts dict') |> snd
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   870
    |> Sign.add_const_constraint (c', SOME ty')
25024
0615bb9955dd added is_class;
wenzelm
parents: 25020
diff changeset
   871
    |> Sign.notation true prmode [(Const (c', ty'), mx)]
25103
haftmann
parents: 25096
diff changeset
   872
    |> register_operation class ((c', (dict', dict'')), NONE)
25096
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   873
    |> Sign.restore_naming thy
24836
dab06e93ec28 intermediate cleanup
haftmann
parents: 24770
diff changeset
   874
  end;
dab06e93ec28 intermediate cleanup
haftmann
parents: 24770
diff changeset
   875
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   876
end;