src/Pure/Isar/class.ML
author haftmann
Fri, 30 Nov 2007 20:13:08 +0100
changeset 25514 4b508bb31a6c
parent 25502 9200b36280c0
child 25518 00d5cc16e891
permissions -rw-r--r--
first working version of instance target
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
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
    10
  (*classes*)
25002
haftmann
parents: 24981
diff changeset
    11
  val class: bstring -> class list -> Element.context_i Locale.element list
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    12
    -> string list -> theory -> string * Proof.context
25002
haftmann
parents: 24981
diff changeset
    13
  val class_cmd: bstring -> xstring list -> Element.context Locale.element list
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    14
    -> xstring list -> theory -> string * Proof.context
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
    15
25311
aa750e3a581c Class.init now similiar to Locale.init
haftmann
parents: 25268
diff changeset
    16
  val init: class -> theory -> Proof.context
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
    17
  val logical_const: string -> Markup.property list
25104
26b9a7db3386 tuned interfaces;
wenzelm
parents: 25103
diff changeset
    18
    -> (string * mixfix) * term -> theory -> theory
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
    19
  val syntactic_const: string -> Syntax.mode -> Markup.property list
25104
26b9a7db3386 tuned interfaces;
wenzelm
parents: 25103
diff changeset
    20
    -> (string * mixfix) * term -> theory -> theory
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
    21
  val refresh_syntax: class -> Proof.context -> Proof.context
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
    22
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    23
  val intro_classes_tac: thm list -> tactic
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    24
  val default_intro_classes_tac: thm list -> tactic
25195
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
    25
  val prove_subclass: class * class -> thm list -> Proof.context
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
    26
    -> theory -> theory
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
    27
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
    28
  val class_prefix: string -> string
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
    29
  val is_class: theory -> class -> bool
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
    30
  val these_params: theory -> sort -> (string * (string * typ)) list
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    31
  val print_classes: theory -> unit
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
    32
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
    33
  (*instances*)
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
    34
  val init_instantiation: arity list -> theory -> local_theory
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
    35
  val instantiation_instance: (local_theory -> local_theory) -> local_theory -> Proof.state
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
    36
  val prove_instantiation_instance: (Proof.context -> tactic) -> local_theory -> local_theory
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
    37
  val conclude_instantiation: local_theory -> local_theory
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
    38
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
    39
  val overloaded_const: string * typ * mixfix -> theory -> term * theory
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
    40
  val overloaded_def: string -> string * term -> theory -> thm * theory
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
    41
  val instantiation_param: Proof.context -> string -> string option
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
    42
  val confirm_declaration: string -> local_theory -> local_theory
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
    43
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
    44
  val unoverload: theory -> thm -> thm
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
    45
  val overload: theory -> thm -> thm
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
    46
  val unoverload_conv: theory -> conv
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
    47
  val overload_conv: theory -> conv
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
    48
  val unoverload_const: theory -> string * typ -> string
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
    49
  val param_of_inst: theory -> string * string -> string
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
    50
  val inst_of_param: theory -> string -> (string * string) option
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
    51
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
    52
  (*old axclass layer*)
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
    53
  val axclass_cmd: bstring * xstring list
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
    54
    -> ((bstring * Attrib.src list) * string list) list
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
    55
    -> theory -> class * theory
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
    56
  val classrel_cmd: xstring * xstring -> theory -> Proof.state
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
    57
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
    58
  (*old instance layer*)
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    59
  val instance_arity: (theory -> theory) -> arity list -> theory -> Proof.state
25502
9200b36280c0 instance command as rudimentary class target
haftmann
parents: 25485
diff changeset
    60
  val instance_arity_cmd: (bstring * xstring list * xstring) list -> theory -> Proof.state
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    61
end;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    62
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    63
structure Class : CLASS =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    64
struct
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    65
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    66
(** auxiliary **)
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    67
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
    68
val classN = "class";
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
    69
val introN = "intro";
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
    70
25002
haftmann
parents: 24981
diff changeset
    71
fun prove_interpretation tac prfx_atts expr inst =
haftmann
parents: 24981
diff changeset
    72
  Locale.interpretation_i I prfx_atts expr inst
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    73
  #> Proof.global_terminal_proof
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    74
      (Method.Basic (K (Method.SIMPLE_METHOD tac), Position.none), NONE)
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    75
  #> ProofContext.theory_of;
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    76
25195
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
    77
fun prove_interpretation_in tac after_qed (name, expr) =
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
    78
  Locale.interpretation_in_locale
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
    79
      (ProofContext.theory after_qed) (name, expr)
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
    80
  #> Proof.global_terminal_proof
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
    81
      (Method.Basic (K (Method.SIMPLE_METHOD tac), Position.none), NONE)
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
    82
  #> ProofContext.theory_of;
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
    83
25020
f1344290e420 (un)overload: full rewrite;
wenzelm
parents: 25010
diff changeset
    84
fun OF_LAST thm1 thm2 = thm1 RSN (Thm.nprems_of thm2, thm2);
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    85
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    86
fun strip_all_ofclass thy sort =
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    87
  let
24847
bc15dcaed517 replaced AxClass.param_tyvarname by Name.aT;
wenzelm
parents: 24836
diff changeset
    88
    val typ = TVar ((Name.aT, 0), sort);
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    89
    fun prem_inclass t =
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    90
      case Logic.strip_imp_prems t
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    91
       of ofcls :: _ => try Logic.dest_inclass ofcls
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    92
        | [] => NONE;
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    93
    fun strip_ofclass class thm =
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    94
      thm OF (fst o AxClass.of_sort thy (typ, [class])) AxClass.cache;
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    95
    fun strip thm = case (prem_inclass o Thm.prop_of) thm
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    96
     of SOME (_, class) => thm |> strip_ofclass class |> strip
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    97
      | NONE => thm;
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    98
  in strip end;
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    99
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   100
fun get_remove_global_constraint c thy =
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   101
  let
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   102
    val ty = Sign.the_const_constraint thy c;
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   103
  in
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   104
    thy
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   105
    |> Sign.add_const_constraint (c, NONE)
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   106
    |> pair (c, Logic.unvarifyT ty)
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   107
  end;
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   108
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   109
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   110
(** primitive axclass and instance commands **)
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   111
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   112
fun axclass_cmd (class, raw_superclasses) raw_specs thy =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   113
  let
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   114
    val ctxt = ProofContext.init thy;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   115
    val superclasses = map (Sign.read_class thy) raw_superclasses;
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   116
    val name_atts = map ((apsnd o map) (Attrib.attribute thy) o fst)
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   117
      raw_specs;
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   118
    val axiomss = ProofContext.read_propp (ctxt, map (map (rpair []) o snd)
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   119
          raw_specs)
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   120
      |> snd
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   121
      |> (map o map) fst;
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   122
  in
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   123
    AxClass.define_class (class, superclasses) []
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   124
      (name_atts ~~ axiomss) thy
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   125
  end;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   126
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   127
local
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   128
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   129
fun gen_instance mk_prop add_thm after_qed insts thy =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   130
  let
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   131
    fun after_qed' results =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   132
      ProofContext.theory ((fold o fold) add_thm results #> after_qed);
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   133
  in
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   134
    thy
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   135
    |> ProofContext.init
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   136
    |> Proof.theorem_i NONE after_qed' ((map (fn t => [(t, [])])
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   137
        o maps (mk_prop thy)) insts)
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   138
  end;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   139
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   140
in
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   141
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   142
val instance_arity =
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   143
  gen_instance (Logic.mk_arities oo Sign.cert_arity) AxClass.add_arity;
25502
9200b36280c0 instance command as rudimentary class target
haftmann
parents: 25485
diff changeset
   144
val instance_arity_cmd =
9200b36280c0 instance command as rudimentary class target
haftmann
parents: 25485
diff changeset
   145
  gen_instance (Logic.mk_arities oo Sign.read_arity) AxClass.add_arity I;
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   146
val classrel =
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   147
  gen_instance (single oo (Logic.mk_classrel oo AxClass.cert_classrel))
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   148
    AxClass.add_classrel I o single;
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   149
val classrel_cmd =
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   150
  gen_instance (single oo (Logic.mk_classrel oo AxClass.read_classrel))
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   151
    AxClass.add_classrel I o single;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   152
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   153
end; (*local*)
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   154
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   155
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   156
(** basic overloading **)
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   157
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   158
(* bookkeeping *)
24304
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   159
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   160
structure InstData = TheoryDataFun
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   161
(
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
   162
  type T = (string * thm) Symtab.table Symtab.table * (string * string) Symtab.table;
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
   163
    (*constant name ~> type constructor ~> (constant name, equation),
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
   164
        constant name ~> (constant name, type constructor)*)
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
   165
  val empty = (Symtab.empty, Symtab.empty);
24304
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   166
  val copy = I;
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   167
  val extend = I;
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
   168
  fun merge _ ((taba1, tabb1), (taba2, tabb2)) =
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
   169
    (Symtab.join (K (Symtab.merge (K true))) (taba1, taba2),
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
   170
      Symtab.merge (K true) (tabb1, tabb2));
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
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   173
val inst_tyco = Option.map fst o try (dest_Type o the_single) oo Sign.const_typargs;
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   174
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   175
fun inst thy (c, tyco) =
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   176
  (the o Symtab.lookup ((the o Symtab.lookup (fst (InstData.get thy))) c)) tyco;
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   177
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   178
val param_of_inst = fst oo inst;
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   179
25020
f1344290e420 (un)overload: full rewrite;
wenzelm
parents: 25010
diff changeset
   180
fun inst_thms thy = (Symtab.fold (Symtab.fold (cons o snd o snd) o snd) o fst)
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   181
  (InstData.get thy) [];
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   182
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   183
val inst_of_param = Symtab.lookup o snd o InstData.get;
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   184
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   185
fun add_inst (c, tyco) inst = (InstData.map o apfst
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   186
      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
   187
  #> (InstData.map o apsnd) (Symtab.update_new (fst inst, (c, tyco)));
24304
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   188
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   189
fun unoverload thy = MetaSimplifier.simplify true (inst_thms thy);
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   190
fun overload thy = MetaSimplifier.simplify true (map Thm.symmetric (inst_thms thy));
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   191
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   192
fun unoverload_conv thy = MetaSimplifier.rewrite true (inst_thms thy);
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   193
fun overload_conv thy = MetaSimplifier.rewrite true (map Thm.symmetric (inst_thms thy));
24304
69d40a562ba4 explicit constants for overloaded definitions
haftmann
parents: 24276
diff changeset
   194
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
   195
fun unoverload_const thy (c_ty as (c, _)) =
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
   196
  case AxClass.class_of_param thy c
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   197
   of SOME class => (case inst_tyco thy c_ty
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   198
       of SOME tyco => try (param_of_inst thy) (c, tyco) |> the_default c
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   199
        | NONE => c)
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
   200
    | NONE => c;
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
   201
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   202
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   203
(* declaration and definition of instances of overloaded constants *)
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   204
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   205
fun primitive_note kind (name, thm) =
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   206
  PureThy.note_thmss_i kind [((name, []), [([thm], [])])]
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   207
  #>> (fn [(_, [thm])] => thm);
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   208
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   209
fun overloaded_const (c, ty, mx) thy =
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   210
  let
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   211
    val _ = if mx <> NoSyn then
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   212
      error ("Illegal mixfix syntax for constant to be instantiated " ^ quote c)
25514
4b508bb31a6c first working version of instance target
haftmann
parents: 25502
diff changeset
   213
      else ();
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   214
    val SOME class = AxClass.class_of_param thy c;
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   215
    val SOME tyco = inst_tyco thy (c, ty);
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   216
    val name_inst = AxClass.instance_name (tyco, class) ^ "_inst";
25502
9200b36280c0 instance command as rudimentary class target
haftmann
parents: 25485
diff changeset
   217
    val c' = NameSpace.base c ^ "_" ^ NameSpace.base tyco;
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   218
    val ty' = Type.strip_sorts ty;
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   219
  in
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   220
    thy
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   221
    |> Sign.sticky_prefix name_inst
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   222
    |> Sign.no_base_names
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   223
    |> Sign.declare_const [] (c', ty', NoSyn)
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   224
    |-> (fn const' as Const (c'', _) => Thm.add_def true
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   225
          (Thm.def_name c', Logic.mk_equals (Const (c, ty'), const'))
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   226
    #>> Thm.varifyT
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   227
    #-> (fn thm => add_inst (c, tyco) (c'', thm)
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   228
    #> primitive_note Thm.internalK (c', thm)
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   229
    #> snd
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   230
    #> Sign.restore_naming thy
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   231
    #> pair (Const (c, ty))))
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   232
  end;
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   233
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   234
fun overloaded_def name (c, t) thy =
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   235
  let
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   236
    val ty = Term.fastype_of t;
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   237
    val SOME tyco = inst_tyco thy (c, ty);
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   238
    val (c', eq) = inst thy (c, tyco);
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   239
    val prop = Logic.mk_equals (Const (c', ty), t);
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   240
    val name' = Thm.def_name_optional
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   241
      (NameSpace.base c ^ "_" ^ NameSpace.base tyco) name;
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   242
  in
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   243
    thy
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   244
    |> Thm.add_def false (name', prop)
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   245
    |>> (fn thm =>  Drule.transitive_thm OF [eq, thm])
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   246
  end;
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   247
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   248
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   249
(** class data **)
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   250
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   251
datatype class_data = ClassData of {
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   252
  consts: (string * string) list
24836
dab06e93ec28 intermediate cleanup
haftmann
parents: 24770
diff changeset
   253
    (*locale parameter ~> constant name*),
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   254
  base_sort: sort,
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   255
  inst: term option list
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   256
    (*canonical interpretation*),
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   257
  morphism: morphism,
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   258
    (*partial morphism of canonical interpretation*)
24657
185502d54c3d fixed wrong syntax treatment in class target
haftmann
parents: 24589
diff changeset
   259
  intro: thm,
185502d54c3d fixed wrong syntax treatment in class target
haftmann
parents: 24589
diff changeset
   260
  defs: thm list,
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   261
  operations: (string * (class * (typ * term))) list
24657
185502d54c3d fixed wrong syntax treatment in class target
haftmann
parents: 24589
diff changeset
   262
};
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   263
24657
185502d54c3d fixed wrong syntax treatment in class target
haftmann
parents: 24589
diff changeset
   264
fun rep_class_data (ClassData d) = d;
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   265
fun mk_class_data ((consts, base_sort, inst, morphism, intro),
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   266
    (defs, operations)) =
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   267
  ClassData { consts = consts, base_sort = base_sort, inst = inst,
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   268
    morphism = morphism, intro = intro, defs = defs,
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   269
    operations = operations };
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   270
fun map_class_data f (ClassData { consts, base_sort, inst, morphism, intro,
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   271
    defs, operations }) =
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   272
  mk_class_data (f ((consts, base_sort, inst, morphism, intro),
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   273
    (defs, operations)));
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   274
fun merge_class_data _ (ClassData { consts = consts,
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   275
    base_sort = base_sort, inst = inst, morphism = morphism, intro = intro,
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   276
    defs = defs1, operations = operations1 },
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   277
  ClassData { consts = _, base_sort = _, inst = _, morphism = _, intro = _,
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   278
    defs = defs2, operations = operations2 }) =
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   279
  mk_class_data ((consts, base_sort, inst, morphism, intro),
24914
95cda5dd58d5 added proper subclass concept; improved class target
haftmann
parents: 24901
diff changeset
   280
    (Thm.merge_thms (defs1, defs2),
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   281
      AList.merge (op =) (K true) (operations1, operations2)));
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   282
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   283
structure ClassData = TheoryDataFun
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   284
(
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   285
  type T = class_data Graph.T
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   286
  val empty = Graph.empty;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   287
  val copy = I;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   288
  val extend = I;
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   289
  fun merge _ = Graph.join merge_class_data;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   290
);
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   291
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   292
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   293
(* queries *)
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   294
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   295
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
   296
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   297
fun the_class_data thy class = case lookup_class_data thy class
25020
f1344290e420 (un)overload: full rewrite;
wenzelm
parents: 25010
diff changeset
   298
 of NONE => error ("Undeclared class " ^ quote class)
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   299
  | SOME data => data;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   300
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   301
val is_class = is_some oo lookup_class_data;
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   302
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   303
val ancestry = Graph.all_succs o ClassData.get;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   304
25002
haftmann
parents: 24981
diff changeset
   305
fun these_params thy =
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   306
  let
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   307
    fun params class =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   308
      let
24930
cc2e0e8c81af renamed AxClass.get_definition to AxClass.get_info (again);
wenzelm
parents: 24920
diff changeset
   309
        val const_typs = (#params o AxClass.get_info thy) class;
24657
185502d54c3d fixed wrong syntax treatment in class target
haftmann
parents: 24589
diff changeset
   310
        val const_names = (#consts o the_class_data thy) class;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   311
      in
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   312
        (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
   313
      end;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   314
  in maps params o ancestry thy end;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   315
24657
185502d54c3d fixed wrong syntax treatment in class target
haftmann
parents: 24589
diff changeset
   316
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
   317
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   318
fun morphism thy = #morphism o the_class_data thy;
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   319
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   320
fun these_intros thy =
24657
185502d54c3d fixed wrong syntax treatment in class target
haftmann
parents: 24589
diff changeset
   321
  Graph.fold (fn (_, (data, _)) => insert Thm.eq_thm ((#intro o rep_class_data) data))
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   322
    (ClassData.get thy) [];
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   323
24836
dab06e93ec28 intermediate cleanup
haftmann
parents: 24770
diff changeset
   324
fun these_operations thy =
dab06e93ec28 intermediate cleanup
haftmann
parents: 24770
diff changeset
   325
  maps (#operations o the_class_data thy) o ancestry thy;
24657
185502d54c3d fixed wrong syntax treatment in class target
haftmann
parents: 24589
diff changeset
   326
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   327
fun print_classes thy =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   328
  let
24920
2a45e400fdad generic Syntax.pretty/string_of operations;
wenzelm
parents: 24914
diff changeset
   329
    val ctxt = ProofContext.init thy;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   330
    val algebra = Sign.classes_of thy;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   331
    val arities =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   332
      Symtab.empty
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   333
      |> Symtab.fold (fn (tyco, arities) => fold (fn (class, _) =>
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   334
           Symtab.map_default (class, []) (insert (op =) tyco)) arities)
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   335
             ((#arities o Sorts.rep_algebra) algebra);
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   336
    val the_arities = these o Symtab.lookup arities;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   337
    fun mk_arity class tyco =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   338
      let
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   339
        val Ss = Sorts.mg_domain algebra tyco [class];
24920
2a45e400fdad generic Syntax.pretty/string_of operations;
wenzelm
parents: 24914
diff changeset
   340
      in Syntax.pretty_arity ctxt (tyco, Ss, [class]) end;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   341
    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
   342
      ^ setmp show_sorts false (Syntax.string_of_typ ctxt o Type.strip_sorts) ty);
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   343
    fun mk_entry class = (Pretty.block o Pretty.fbreaks o map_filter I) [
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   344
      (SOME o Pretty.str) ("class " ^ Sign.extern_class thy class ^ ":"),
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   345
      (SOME o Pretty.block) [Pretty.str "supersort: ",
24920
2a45e400fdad generic Syntax.pretty/string_of operations;
wenzelm
parents: 24914
diff changeset
   346
        (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
   347
      if is_class thy class then (SOME o Pretty.str)
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   348
        ("locale: " ^ Locale.extern thy class) else NONE,
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   349
      ((fn [] => NONE | ps => (SOME o Pretty.block o Pretty.fbreaks)
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   350
          (Pretty.str "parameters:" :: ps)) o map mk_param
24930
cc2e0e8c81af renamed AxClass.get_definition to AxClass.get_info (again);
wenzelm
parents: 24920
diff changeset
   351
        o these o Option.map #params o try (AxClass.get_info thy)) class,
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   352
      (SOME o Pretty.block o Pretty.breaks) [
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   353
        Pretty.str "instances:",
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   354
        Pretty.list "" "" (map (mk_arity class) (the_arities class))
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   355
      ]
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   356
    ]
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   357
  in
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   358
    (Pretty.writeln o Pretty.chunks o separate (Pretty.str "")
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   359
      o map mk_entry o Sorts.all_classes) algebra
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   360
  end;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   361
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   362
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   363
(* updaters *)
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   364
25163
haftmann
parents: 25146
diff changeset
   365
fun add_class_data ((class, superclasses), (cs, base_sort, inst, phi, intro)) thy =
25002
haftmann
parents: 24981
diff changeset
   366
  let
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   367
    val operations = map (fn (v_ty as (_, ty), (c, _)) =>
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   368
      (c, (class, (ty, Free v_ty)))) cs;
25002
haftmann
parents: 24981
diff changeset
   369
    val cs = (map o pairself) fst cs;
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   370
    val add_class = Graph.new_node (class,
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   371
        mk_class_data ((cs, base_sort, map (SOME o Const) inst, phi, intro), ([], operations)))
25002
haftmann
parents: 24981
diff changeset
   372
      #> fold (curry Graph.add_edge class) superclasses;
haftmann
parents: 24981
diff changeset
   373
  in
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   374
    ClassData.map add_class thy
25002
haftmann
parents: 24981
diff changeset
   375
  end;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   376
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   377
fun register_operation class (c, (t, some_def)) thy =
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   378
  let
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   379
    val base_sort = (#base_sort o the_class_data thy) class;
25239
3d6abdd10382 handling of notation in class target
haftmann
parents: 25210
diff changeset
   380
    val prep_typ = map_atyps
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   381
      (fn TVar (vi as (v, _), sort) => if Name.aT = v
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   382
        then TFree (v, base_sort) else TVar (vi, sort));
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   383
    val t' = map_types prep_typ t;
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   384
    val ty' = Term.fastype_of t';
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   385
  in
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   386
    thy
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   387
    |> (ClassData.map o Graph.map_node class o map_class_data o apsnd)
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   388
      (fn (defs, operations) =>
25096
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   389
        (fold cons (the_list some_def) defs,
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   390
          (c, (class, (ty', t'))) :: operations))
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   391
  end;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   392
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   393
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   394
(** rule calculation, tactics and methods **)
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   395
25024
0615bb9955dd added is_class;
wenzelm
parents: 25020
diff changeset
   396
val class_prefix = Logic.const_of_class o Sign.base_name;
0615bb9955dd added is_class;
wenzelm
parents: 25020
diff changeset
   397
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   398
fun calculate_morphism class cs =
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   399
  let
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   400
    val subst_typ = Term.map_type_tfree (fn var as (v, sort) =>
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   401
      if v = Name.aT then TVar ((v, 0), [class]) else TVar ((v, 0), sort));
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   402
    fun subst_aterm (t as Free (v, ty)) = (case AList.lookup (op =) cs v
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   403
         of SOME (c, _) => Const (c, ty)
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   404
          | NONE => t)
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   405
      | subst_aterm t = t;
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   406
    val subst_term = map_aterms subst_aterm #> map_types subst_typ;
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   407
  in
25209
haftmann
parents: 25195
diff changeset
   408
    Morphism.term_morphism subst_term
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   409
    $> Morphism.typ_morphism subst_typ
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   410
  end;
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   411
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   412
fun class_intro thy class sups =
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   413
  let
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   414
    fun class_elim class =
25020
f1344290e420 (un)overload: full rewrite;
wenzelm
parents: 25010
diff changeset
   415
      case (#axioms o AxClass.get_info thy) class
f1344290e420 (un)overload: full rewrite;
wenzelm
parents: 25010
diff changeset
   416
       of [thm] => SOME (Drule.unconstrainTs thm)
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   417
        | [] => NONE;
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   418
    val pred_intro = case Locale.intros thy class
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   419
     of ([ax_intro], [intro]) => intro |> OF_LAST ax_intro |> SOME
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   420
      | ([intro], []) => SOME intro
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   421
      | ([], [intro]) => SOME intro
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   422
      | _ => NONE;
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   423
    val pred_intro' = pred_intro
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   424
      |> 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
   425
    val class_intro = (#intro o AxClass.get_info thy) class;
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   426
    val raw_intro = case pred_intro'
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   427
     of SOME pred_intro => class_intro |> OF_LAST pred_intro
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   428
      | NONE => class_intro;
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   429
    val sort = Sign.super_classes thy class;
24847
bc15dcaed517 replaced AxClass.param_tyvarname by Name.aT;
wenzelm
parents: 24836
diff changeset
   430
    val typ = TVar ((Name.aT, 0), sort);
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   431
    val defs = these_defs thy sups;
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   432
  in
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   433
    raw_intro
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   434
    |> Drule.instantiate' [SOME (Thm.ctyp_of thy typ)] []
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   435
    |> strip_all_ofclass thy sort
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   436
    |> Thm.strip_shyps
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   437
    |> MetaSimplifier.rewrite_rule defs
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   438
    |> Drule.unconstrainTs
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   439
  end;
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   440
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   441
fun class_interpretation class facts defs thy =
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   442
  let
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   443
    val params = these_params thy [class];
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   444
    val inst = (#inst o the_class_data thy) class;
25020
f1344290e420 (un)overload: full rewrite;
wenzelm
parents: 25010
diff changeset
   445
    val tac = ALLGOALS (ProofContext.fact_tac facts);
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   446
    val prfx = class_prefix class;
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   447
  in
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   448
    thy
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   449
    |> 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
   450
    ||> prove_interpretation tac ((false, prfx), []) (Locale.Locale class)
ba43514068fd Interpretation equations may have name and/or attribute.
ballarin
parents: 25083
diff changeset
   451
          (inst, map (fn def => (("", []), def)) defs)
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   452
    |-> (fn cs => fold (Sign.add_const_constraint o apsnd SOME) cs)
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   453
  end;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   454
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   455
fun intro_classes_tac facts st =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   456
  let
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   457
    val thy = Thm.theory_of_thm st;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   458
    val classes = Sign.all_classes thy;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   459
    val class_trivs = map (Thm.class_triv thy) classes;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   460
    val class_intros = these_intros thy;
24930
cc2e0e8c81af renamed AxClass.get_definition to AxClass.get_info (again);
wenzelm
parents: 24920
diff changeset
   461
    val axclass_intros = map_filter (try (#intro o AxClass.get_info thy)) classes;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   462
  in
25268
58146cb7bf44 more precise treatment of prove_subclass
haftmann
parents: 25239
diff changeset
   463
    Method.intros_tac (class_trivs @ class_intros @ axclass_intros) facts st
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   464
  end;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   465
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   466
fun default_intro_classes_tac [] = intro_classes_tac []
24930
cc2e0e8c81af renamed AxClass.get_definition to AxClass.get_info (again);
wenzelm
parents: 24920
diff changeset
   467
  | default_intro_classes_tac _ = no_tac;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   468
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   469
fun default_tac rules ctxt facts =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   470
  HEADGOAL (Method.some_rule_tac rules ctxt facts) ORELSE
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   471
    default_intro_classes_tac facts;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   472
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   473
val _ = Context.add_setup (Method.add_methods
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   474
 [("intro_classes", Method.no_args (Method.METHOD intro_classes_tac),
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   475
    "back-chain introduction rules of classes"),
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   476
  ("default", Method.thms_ctxt_args (Method.METHOD oo default_tac),
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   477
    "apply some intro/elim rule")]);
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   478
25195
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   479
fun subclass_rule thy (sub, sup) =
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   480
  let
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   481
    val ctxt = Locale.init sub thy;
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   482
    val ctxt_thy = ProofContext.init thy;
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   483
    val props =
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   484
      Locale.global_asms_of thy sup
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   485
      |> maps snd
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   486
      |> map (ObjectLogic.ensure_propT thy);
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   487
    fun tac { prems, context } =
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   488
      Locale.intro_locales_tac true context prems
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   489
        ORELSE ALLGOALS assume_tac;
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   490
  in
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   491
    Goal.prove_multi ctxt [] [] props tac
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   492
    |> map (Assumption.export false ctxt ctxt_thy)
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   493
    |> Variable.export ctxt ctxt_thy
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   494
  end;
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   495
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   496
fun prove_single_subclass (sub, sup) thms ctxt thy =
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   497
  let
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   498
    val ctxt_thy = ProofContext.init thy;
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   499
    val subclass_rule = Conjunction.intr_balanced thms
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   500
      |> Assumption.export false ctxt ctxt_thy
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   501
      |> singleton (Variable.export ctxt ctxt_thy);
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   502
    val sub_inst = Thm.ctyp_of thy (TVar ((Name.aT, 0), [sub]));
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   503
    val sub_ax = #axioms (AxClass.get_info thy sub);
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   504
    val classrel =
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   505
      #intro (AxClass.get_info thy sup)
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   506
      |> Drule.instantiate' [SOME sub_inst] []
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   507
      |> OF_LAST (subclass_rule OF sub_ax)
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   508
      |> strip_all_ofclass thy (Sign.super_classes thy sup)
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   509
      |> Thm.strip_shyps
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   510
  in
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   511
    thy
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   512
    |> AxClass.add_classrel classrel
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   513
    |> prove_interpretation_in (ALLGOALS (ProofContext.fact_tac thms))
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   514
         I (sub, Locale.Locale sup)
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   515
    |> ClassData.map (Graph.add_edge (sub, sup))
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   516
  end;
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   517
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   518
fun prove_subclass (sub, sup) thms ctxt thy =
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   519
  let
25268
58146cb7bf44 more precise treatment of prove_subclass
haftmann
parents: 25239
diff changeset
   520
    val classes = ClassData.get thy;
58146cb7bf44 more precise treatment of prove_subclass
haftmann
parents: 25239
diff changeset
   521
    val is_sup = not o null o curry (Graph.irreducible_paths classes) sub;
58146cb7bf44 more precise treatment of prove_subclass
haftmann
parents: 25239
diff changeset
   522
    val supclasses = Graph.all_succs classes [sup] |> filter_out is_sup;
25195
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   523
    fun transform sup' = subclass_rule thy (sup, sup') |> map (fn thm => thm OF thms);
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   524
  in
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   525
    thy
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   526
    |> fold_rev (fn sup' => prove_single_subclass (sub, sup')
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   527
         (transform sup') ctxt) supclasses
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   528
 end;
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   529
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   530
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   531
(** classes and class target **)
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   532
25002
haftmann
parents: 24981
diff changeset
   533
(* class context syntax *)
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   534
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   535
structure ClassSyntax = ProofDataFun(
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   536
  type T = {
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   537
    local_constraints: (string * typ) list,
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   538
    global_constraints: (string * typ) list,
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   539
    base_sort: sort,
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   540
    operations: (string * (typ * term)) list,
25195
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   541
    unchecks: (term * term) list,
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   542
    passed: bool
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   543
  };
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   544
  fun init _ = {
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   545
    local_constraints = [],
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   546
    global_constraints = [],
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   547
    base_sort = [],
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   548
    operations = [],
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   549
    unchecks = [],
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   550
    passed = true
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   551
  };;
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   552
);
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   553
25344
00c2179db769 synchronize_syntax: improved declare_const (still inactive);
wenzelm
parents: 25326
diff changeset
   554
fun synchronize_syntax sups base_sort ctxt =
24914
95cda5dd58d5 added proper subclass concept; improved class target
haftmann
parents: 24901
diff changeset
   555
  let
25344
00c2179db769 synchronize_syntax: improved declare_const (still inactive);
wenzelm
parents: 25326
diff changeset
   556
    val thy = ProofContext.theory_of ctxt;
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   557
    fun subst_class_typ sort = map_atyps
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   558
      (fn TFree _ => TVar ((Name.aT, 0), sort) | ty' => ty');
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   559
    val operations = these_operations thy sups;
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   560
    val local_constraints =
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   561
      (map o apsnd) (subst_class_typ base_sort o fst o snd) operations;
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   562
    val global_constraints =
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   563
      (map o apsnd) (fn (class, (ty, _)) => subst_class_typ [class] ty) operations;
25318
c8352b38d47d synchronize_syntax: declare operations within the local scope of fixes/consts
wenzelm
parents: 25311
diff changeset
   564
    fun declare_const (c, _) =
c8352b38d47d synchronize_syntax: declare operations within the local scope of fixes/consts
wenzelm
parents: 25311
diff changeset
   565
      let val b = Sign.base_name c
25344
00c2179db769 synchronize_syntax: improved declare_const (still inactive);
wenzelm
parents: 25326
diff changeset
   566
      in Sign.intern_const thy b = c ? Variable.declare_const (b, c) end;
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   567
    val unchecks = map (fn (c, (_, (ty, t))) => (t, Const (c, ty))) operations;
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   568
  in
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   569
    ctxt
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   570
    |> fold declare_const local_constraints
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   571
    |> fold (ProofContext.add_const_constraint o apsnd SOME) local_constraints
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   572
    |> ClassSyntax.put {
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   573
        local_constraints = local_constraints,
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   574
        global_constraints = global_constraints,
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   575
        base_sort = base_sort,
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   576
        operations = (map o apsnd) snd operations,
25195
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   577
        unchecks = unchecks,
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   578
        passed = false
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   579
      }
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   580
  end;
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   581
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   582
fun refresh_syntax class ctxt =
25002
haftmann
parents: 24981
diff changeset
   583
  let
haftmann
parents: 24981
diff changeset
   584
    val thy = ProofContext.theory_of ctxt;
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   585
    val base_sort = (#base_sort o the_class_data thy) class;
25344
00c2179db769 synchronize_syntax: improved declare_const (still inactive);
wenzelm
parents: 25326
diff changeset
   586
  in synchronize_syntax [class] base_sort ctxt end;
24914
95cda5dd58d5 added proper subclass concept; improved class target
haftmann
parents: 24901
diff changeset
   587
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   588
val mark_passed = ClassSyntax.map
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   589
  (fn { local_constraints, global_constraints, base_sort, operations, unchecks, passed } =>
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   590
    { local_constraints = local_constraints, global_constraints = global_constraints,
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   591
      base_sort = base_sort, operations = operations, unchecks = unchecks, passed = true });
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   592
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   593
fun sort_term_check ts ctxt =
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   594
  let
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   595
    val { local_constraints, global_constraints, base_sort, operations, passed, ... } =
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   596
      ClassSyntax.get ctxt;
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   597
    fun check_improve (Const (c, ty)) = (case AList.lookup (op =) local_constraints c
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   598
         of SOME ty0 => (case try (Type.raw_match (ty0, ty)) Vartab.empty
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   599
             of SOME tyenv => (case Vartab.lookup tyenv (Name.aT, 0)
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   600
                 of SOME (_, TVar (tvar as (vi, _))) =>
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   601
                      if TypeInfer.is_param vi then cons tvar else I
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   602
                  | _ => I)
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   603
              | NONE => I)
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   604
          | NONE => I)
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   605
      | check_improve _ = I;
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   606
    val improvements = (fold o fold_aterms) check_improve ts [];
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   607
    val ts' = (map o map_types o map_atyps) (fn ty as TVar tvar =>
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   608
        if member (op =) improvements tvar
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   609
          then TFree (Name.aT, base_sort) else ty | ty => ty) ts;
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   610
    fun check t0 = Envir.expand_term (fn Const (c, ty) => (case AList.lookup (op =) operations c
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   611
         of SOME (ty0, t) =>
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   612
              if Type.typ_instance (ProofContext.tsig_of ctxt) (ty, ty0)
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   613
              then SOME (ty0, check t) else NONE
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   614
          | NONE => NONE)
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   615
      | _ => NONE) t0;
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   616
    val ts'' = map check ts';
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   617
  in if eq_list (op aconv) (ts, ts'') andalso passed then NONE
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   618
  else
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   619
    ctxt
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   620
    |> fold (ProofContext.add_const_constraint o apsnd SOME) global_constraints
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   621
    |> mark_passed
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   622
    |> pair ts''
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   623
    |> SOME
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   624
  end;
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   625
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   626
fun sort_term_uncheck ts ctxt =
25002
haftmann
parents: 24981
diff changeset
   627
  let
haftmann
parents: 24981
diff changeset
   628
    val thy = ProofContext.theory_of ctxt;
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   629
    val unchecks = (#unchecks o ClassSyntax.get) ctxt;
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   630
    val ts' = map (Pattern.rewrite_term thy unchecks []) ts;
25060
17c313217998 Syntax.(un)check: explicit result option;
wenzelm
parents: 25038
diff changeset
   631
  in if eq_list (op aconv) (ts, ts') then NONE else SOME (ts', ctxt) end;
25002
haftmann
parents: 24981
diff changeset
   632
25344
00c2179db769 synchronize_syntax: improved declare_const (still inactive);
wenzelm
parents: 25326
diff changeset
   633
fun init_ctxt sups base_sort ctxt =
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   634
  ctxt
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   635
  |> Variable.declare_term
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   636
      (Logic.mk_type (TFree (Name.aT, base_sort)))
25344
00c2179db769 synchronize_syntax: improved declare_const (still inactive);
wenzelm
parents: 25326
diff changeset
   637
  |> synchronize_syntax sups base_sort
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   638
  |> Context.proof_map (
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   639
      Syntax.add_term_check 0 "class" sort_term_check
25103
haftmann
parents: 25096
diff changeset
   640
      #> 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
   641
25311
aa750e3a581c Class.init now similiar to Locale.init
haftmann
parents: 25268
diff changeset
   642
fun init class thy =
aa750e3a581c Class.init now similiar to Locale.init
haftmann
parents: 25268
diff changeset
   643
  thy
aa750e3a581c Class.init now similiar to Locale.init
haftmann
parents: 25268
diff changeset
   644
  |> Locale.init class
25344
00c2179db769 synchronize_syntax: improved declare_const (still inactive);
wenzelm
parents: 25326
diff changeset
   645
  |> init_ctxt [class] ((#base_sort o the_class_data thy) class);
24914
95cda5dd58d5 added proper subclass concept; improved class target
haftmann
parents: 24901
diff changeset
   646
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   647
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   648
(* class definition *)
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   649
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   650
local
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   651
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   652
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
   653
  let
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   654
    val supclasses = map (prep_class thy) raw_supclasses;
25209
haftmann
parents: 25195
diff changeset
   655
    val sups = filter (is_class thy) supclasses;
haftmann
parents: 25195
diff changeset
   656
    fun the_base_sort class = lookup_class_data thy class
haftmann
parents: 25195
diff changeset
   657
      |> Option.map #base_sort
haftmann
parents: 25195
diff changeset
   658
      |> the_default [class];
haftmann
parents: 25195
diff changeset
   659
    val base_sort = Sign.minimize_sort thy (maps the_base_sort supclasses);
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   660
    val supsort = Sign.minimize_sort thy supclasses;
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   661
    val suplocales = map Locale.Locale sups;
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   662
    val (raw_elems, includes) = fold_rev (fn Locale.Elem e => apfst (cons e)
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   663
      | Locale.Expr i => apsnd (cons (prep_expr thy i))) raw_includes_elems ([], []);
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   664
    val supexpr = Locale.Merge suplocales;
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   665
    val supparams = (map fst o Locale.parameters_of_expr thy) supexpr;
25002
haftmann
parents: 24981
diff changeset
   666
    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
   667
      (map fst supparams);
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   668
    val mergeexpr = Locale.Merge (suplocales @ includes);
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   669
    val constrain = Element.Constrains ((map o apsnd o map_atyps)
24847
bc15dcaed517 replaced AxClass.param_tyvarname by Name.aT;
wenzelm
parents: 24836
diff changeset
   670
      (fn TFree (_, sort) => TFree (Name.aT, sort)) supparams);
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   671
  in
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   672
    ProofContext.init thy
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   673
    |> Locale.cert_expr supexpr [constrain]
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   674
    |> snd
25344
00c2179db769 synchronize_syntax: improved declare_const (still inactive);
wenzelm
parents: 25326
diff changeset
   675
    |> init_ctxt sups base_sort
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   676
    |> process_expr Locale.empty raw_elems
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   677
    |> fst
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   678
    |> (fn elems => ((((sups, supconsts), (supsort, base_sort, mergeexpr)),
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   679
          (*FIXME*) if null includes then constrain :: elems else elems)))
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   680
  end;
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   681
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   682
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
   683
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
   684
24968
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   685
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
   686
  let
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   687
    val superclasses = map (Sign.certify_class thy) raw_superclasses;
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   688
    val consts = (map o apfst o apsnd) (Sign.certify_typ thy) raw_consts;
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   689
    fun add_const ((c, ty), syn) =
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   690
      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
   691
    fun mk_axioms cs thy =
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   692
      raw_dep_axioms thy cs
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   693
      |> (map o apsnd o map) (Sign.cert_prop thy)
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   694
      |> rpair thy;
25002
haftmann
parents: 24981
diff changeset
   695
    fun constrain_typs class = (map o apsnd o Term.map_type_tfree)
haftmann
parents: 24981
diff changeset
   696
      (fn (v, _) => TFree (v, [class]))
24968
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   697
  in
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   698
    thy
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   699
    |> Sign.add_path (Logic.const_of_class name)
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   700
    |> fold_map add_const consts
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   701
    ||> Sign.restore_naming thy
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   702
    |-> (fn cs => mk_axioms cs
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   703
    #-> (fn axioms_prop => AxClass.define_class (name, superclasses)
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   704
           (map fst cs @ other_consts) axioms_prop
25002
haftmann
parents: 24981
diff changeset
   705
    #-> (fn class => `(fn _ => constrain_typs class cs)
haftmann
parents: 24981
diff changeset
   706
    #-> (fn cs' => `(fn thy => AxClass.get_info thy class)
haftmann
parents: 24981
diff changeset
   707
    #-> (fn {axioms, ...} => fold (Sign.add_const_constraint o apsnd SOME) cs'
haftmann
parents: 24981
diff changeset
   708
    #> pair (class, (cs', axioms)))))))
24968
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   709
  end;
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   710
25002
haftmann
parents: 24981
diff changeset
   711
fun gen_class prep_spec prep_param bname
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   712
    raw_supclasses raw_includes_elems raw_other_consts thy =
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   713
  let
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   714
    val class = Sign.full_name thy bname;
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   715
    val (((sups, supconsts), (supsort, base_sort, mergeexpr)), elems_syn) =
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   716
      prep_spec thy raw_supclasses raw_includes_elems;
24968
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   717
    val other_consts = map (tap (Sign.the_const_type thy) o prep_param thy) raw_other_consts;
25163
haftmann
parents: 25146
diff changeset
   718
    fun mk_inst class cs =
haftmann
parents: 25146
diff changeset
   719
      (map o apsnd o Term.map_type_tfree) (fn (v, _) => TFree (v, [class])) cs;
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   720
    fun fork_syntax (Element.Fixes xs) =
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   721
          fold_map (fn (c, ty, syn) => cons (c, syn) #> pair (c, ty, NoSyn)) xs
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   722
          #>> Element.Fixes
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   723
      | fork_syntax x = pair x;
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   724
    val (elems, global_syn) = fold_map fork_syntax elems_syn [];
25326
e417a7236125 refined Variable.declare_const;
wenzelm
parents: 25318
diff changeset
   725
    fun globalize (c, ty) =
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   726
      ((c, Term.map_type_tfree (K (TFree (Name.aT, base_sort))) ty),
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   727
        (the_default NoSyn o AList.lookup (op =) global_syn) c);
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   728
    fun extract_params thy =
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   729
      let
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   730
        val params = map fst (Locale.parameters_of thy class);
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   731
      in
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   732
        (params, (map globalize o snd o chop (length supconsts)) params)
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   733
      end;
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   734
    fun extract_assumes params thy cs =
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   735
      let
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   736
        val consts = supconsts @ (map (fst o fst) params ~~ cs);
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   737
        fun subst (Free (c, ty)) =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   738
              Const ((fst o the o AList.lookup (op =) consts) c, ty)
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   739
          | subst t = t;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   740
        fun prep_asm ((name, atts), ts) =
25024
0615bb9955dd added is_class;
wenzelm
parents: 25020
diff changeset
   741
          ((Sign.base_name name, map (Attrib.attribute_i thy) atts),
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   742
            (map o map_aterms) subst ts);
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   743
      in
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   744
        Locale.global_asms_of thy class
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   745
        |> map prep_asm
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   746
      end;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   747
  in
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   748
    thy
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   749
    |> Locale.add_locale_i (SOME "") bname mergeexpr elems
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   750
    |> snd
25311
aa750e3a581c Class.init now similiar to Locale.init
haftmann
parents: 25268
diff changeset
   751
    |> ProofContext.theory_of
aa750e3a581c Class.init now similiar to Locale.init
haftmann
parents: 25268
diff changeset
   752
    |> `extract_params
aa750e3a581c Class.init now similiar to Locale.init
haftmann
parents: 25268
diff changeset
   753
    |-> (fn (all_params, params) =>
24968
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   754
        define_class_params (bname, supsort) params
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   755
          (extract_assumes params) other_consts
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   756
      #-> (fn (_, (consts, axioms)) =>
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   757
        `(fn thy => class_intro thy class sups)
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   758
      #-> (fn class_intro =>
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   759
        PureThy.note_thmss_qualified "" (NameSpace.append class classN)
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   760
          [((introN, []), [([class_intro], [])])]
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   761
      #-> (fn [(_, [class_intro])] =>
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   762
        add_class_data ((class, sups),
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   763
          (map fst params ~~ consts, base_sort,
25163
haftmann
parents: 25146
diff changeset
   764
            mk_inst class (map snd supconsts @ consts),
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   765
              calculate_morphism class (supconsts @ (map (fst o fst) params ~~ consts)), class_intro))
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   766
      #> class_interpretation class axioms []
25311
aa750e3a581c Class.init now similiar to Locale.init
haftmann
parents: 25268
diff changeset
   767
      ))))
25268
58146cb7bf44 more precise treatment of prove_subclass
haftmann
parents: 25239
diff changeset
   768
    |> init class
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   769
    |> pair class
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   770
  end;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   771
25326
e417a7236125 refined Variable.declare_const;
wenzelm
parents: 25318
diff changeset
   772
fun read_const thy = #1 o Term.dest_Const o ProofContext.read_const (ProofContext.init thy);
e417a7236125 refined Variable.declare_const;
wenzelm
parents: 25318
diff changeset
   773
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   774
in
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   775
25326
e417a7236125 refined Variable.declare_const;
wenzelm
parents: 25318
diff changeset
   776
val class_cmd = gen_class read_class_spec read_const;
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   777
val class = gen_class check_class_spec (K I);
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   778
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   779
end; (*local*)
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   780
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   781
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   782
(* class target *)
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   783
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   784
fun logical_const class pos ((c, mx), dict) thy =
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   785
  let
25024
0615bb9955dd added is_class;
wenzelm
parents: 25020
diff changeset
   786
    val prfx = class_prefix class;
0615bb9955dd added is_class;
wenzelm
parents: 25020
diff changeset
   787
    val thy' = thy |> Sign.add_path prfx;
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   788
    val phi = morphism thy' class;
25024
0615bb9955dd added is_class;
wenzelm
parents: 25020
diff changeset
   789
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   790
    val c' = Sign.full_name thy' c;
25239
3d6abdd10382 handling of notation in class target
haftmann
parents: 25210
diff changeset
   791
    val dict' = Morphism.term phi dict;
3d6abdd10382 handling of notation in class target
haftmann
parents: 25210
diff changeset
   792
    val dict_def = map_types Logic.unvarifyT dict';
3d6abdd10382 handling of notation in class target
haftmann
parents: 25210
diff changeset
   793
    val ty' = Term.fastype_of dict_def;
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   794
    val ty'' = Type.strip_sorts ty';
25239
3d6abdd10382 handling of notation in class target
haftmann
parents: 25210
diff changeset
   795
    val def_eq = Logic.mk_equals (Const (c', ty'), dict_def);
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   796
  in
25024
0615bb9955dd added is_class;
wenzelm
parents: 25020
diff changeset
   797
    thy'
25096
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   798
    |> Sign.declare_const pos (c, ty'', mx) |> snd
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   799
    |> Thm.add_def false (c, def_eq)
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   800
    |>> Thm.symmetric
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   801
    |-> (fn def => class_interpretation class [def] [Thm.prop_of def]
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   802
          #> register_operation class (c', (dict', SOME (Thm.varifyT def))))
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   803
    |> Sign.restore_naming thy
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   804
    |> Sign.add_const_constraint (c', SOME ty')
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   805
  end;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   806
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   807
fun syntactic_const class prmode pos ((c, mx), rhs) thy =
24836
dab06e93ec28 intermediate cleanup
haftmann
parents: 24770
diff changeset
   808
  let
25024
0615bb9955dd added is_class;
wenzelm
parents: 25020
diff changeset
   809
    val prfx = class_prefix class;
25096
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
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;
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   812
25096
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   813
    val c' = Sign.full_name thy' c;
25146
c2a41f31cacb tuned abbreviations in class context
haftmann
parents: 25104
diff changeset
   814
    val rews = map (Logic.dest_equals o Thm.prop_of) (these_defs thy' [class])
c2a41f31cacb tuned abbreviations in class context
haftmann
parents: 25104
diff changeset
   815
    val rhs' = (Pattern.rewrite_term thy rews [] o Morphism.term phi) rhs;
25239
3d6abdd10382 handling of notation in class target
haftmann
parents: 25210
diff changeset
   816
    val ty' = Logic.unvarifyT (Term.fastype_of rhs');
24836
dab06e93ec28 intermediate cleanup
haftmann
parents: 24770
diff changeset
   817
  in
25096
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   818
    thy'
25146
c2a41f31cacb tuned abbreviations in class context
haftmann
parents: 25104
diff changeset
   819
    |> Sign.add_abbrev (#1 prmode) pos (c, map_types Type.strip_sorts rhs') |> snd
25096
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   820
    |> Sign.add_const_constraint (c', SOME ty')
25024
0615bb9955dd added is_class;
wenzelm
parents: 25020
diff changeset
   821
    |> Sign.notation true prmode [(Const (c', ty'), mx)]
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   822
    |> register_operation class (c', (rhs', NONE))
25096
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   823
    |> Sign.restore_naming thy
24836
dab06e93ec28 intermediate cleanup
haftmann
parents: 24770
diff changeset
   824
  end;
dab06e93ec28 intermediate cleanup
haftmann
parents: 24770
diff changeset
   825
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   826
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   827
(** instantiation target **)
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   828
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   829
(* bookkeeping *)
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   830
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   831
datatype instantiation = Instantiation of {
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   832
  arities: arity list,
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   833
  params: ((string * string) * (string * typ)) list
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   834
}
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   835
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   836
structure Instantiation = ProofDataFun
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   837
(
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   838
  type T = instantiation
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   839
  fun init _ = Instantiation { arities = [], params = [] };
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   840
);
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   841
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   842
fun mk_instantiation (arities, params) =
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   843
  Instantiation { arities = arities, params = params };
25514
4b508bb31a6c first working version of instance target
haftmann
parents: 25502
diff changeset
   844
fun get_instantiation lthy = case Instantiation.get (LocalTheory.target_of lthy)
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   845
 of Instantiation data => data;
25514
4b508bb31a6c first working version of instance target
haftmann
parents: 25502
diff changeset
   846
fun map_instantiation f = (LocalTheory.target o Instantiation.map)
4b508bb31a6c first working version of instance target
haftmann
parents: 25502
diff changeset
   847
  (fn Instantiation { arities, params } => mk_instantiation (f (arities, params)));
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   848
25514
4b508bb31a6c first working version of instance target
haftmann
parents: 25502
diff changeset
   849
fun the_instantiation lthy = case get_instantiation lthy
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   850
 of { arities = [], ... } => error "No instantiation target"
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   851
  | data => data;
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   852
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   853
val instantiation_params = #params o get_instantiation;
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   854
25514
4b508bb31a6c first working version of instance target
haftmann
parents: 25502
diff changeset
   855
fun instantiation_param lthy v = instantiation_params lthy
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   856
  |> find_first (fn (_, (v', _)) => v = v')
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   857
  |> Option.map (fst o fst);
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   858
25514
4b508bb31a6c first working version of instance target
haftmann
parents: 25502
diff changeset
   859
fun confirm_declaration c = (map_instantiation o apsnd)
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   860
  (filter_out (fn (_, (c', _)) => c' = c));
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   861
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   862
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   863
(* syntax *)
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   864
25514
4b508bb31a6c first working version of instance target
haftmann
parents: 25502
diff changeset
   865
fun inst_term_check ts lthy =
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   866
  let
25514
4b508bb31a6c first working version of instance target
haftmann
parents: 25502
diff changeset
   867
    val params = instantiation_params lthy;
4b508bb31a6c first working version of instance target
haftmann
parents: 25502
diff changeset
   868
    val tsig = ProofContext.tsig_of lthy;
4b508bb31a6c first working version of instance target
haftmann
parents: 25502
diff changeset
   869
    val thy = ProofContext.theory_of lthy;
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   870
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   871
    fun check_improve (Const (c, ty)) = (case inst_tyco thy (c, ty)
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   872
         of SOME tyco => (case AList.lookup (op =) params (c, tyco)
25502
9200b36280c0 instance command as rudimentary class target
haftmann
parents: 25485
diff changeset
   873
             of SOME (_, ty') => perhaps (try (Type.typ_match tsig (ty, ty')))
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   874
              | NONE => I)
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   875
          | NONE => I)
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   876
      | check_improve _ = I;
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   877
    val improvement = (fold o fold_aterms) check_improve ts Vartab.empty;
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   878
    val ts' = (map o map_types) (Envir.typ_subst_TVars improvement) ts;
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   879
    val ts'' = (map o map_aterms) (fn t as Const (c, ty) => (case inst_tyco thy (c, ty)
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   880
         of SOME tyco => (case AList.lookup (op =) params (c, tyco)
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   881
             of SOME v_ty => Free v_ty
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   882
              | NONE => t)
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   883
          | NONE => t)
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   884
      | t => t) ts';
25514
4b508bb31a6c first working version of instance target
haftmann
parents: 25502
diff changeset
   885
  in if eq_list (op aconv) (ts, ts'') then NONE else SOME (ts'', lthy) end;
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   886
25514
4b508bb31a6c first working version of instance target
haftmann
parents: 25502
diff changeset
   887
fun inst_term_uncheck ts lthy =
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   888
  let
25514
4b508bb31a6c first working version of instance target
haftmann
parents: 25502
diff changeset
   889
    val params = instantiation_params lthy;
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   890
    val ts' = (map o map_aterms) (fn t as Free (v, ty) =>
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   891
       (case get_first (fn ((c, _), (v', _)) => if v = v' then SOME c else NONE) params
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   892
         of SOME c => Const (c, ty)
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   893
          | NONE => t)
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   894
      | t => t) ts;
25514
4b508bb31a6c first working version of instance target
haftmann
parents: 25502
diff changeset
   895
  in if eq_list (op aconv) (ts, ts') then NONE else SOME (ts', lthy) end;
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   896
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   897
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   898
(* target *)
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   899
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   900
val sanatize_name = (*FIXME*)
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   901
  let
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   902
    fun is_valid s = Symbol.is_ascii_letter s orelse Symbol.is_ascii_digit s orelse s = "'";
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   903
    val is_junk = not o is_valid andf Symbol.is_regular;
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   904
    val junk = Scan.many is_junk;
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   905
    val scan_valids = Symbol.scanner "Malformed input"
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   906
      ((junk |--
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   907
        (Scan.optional (Scan.one Symbol.is_ascii_letter) "x" ^^ (Scan.many is_valid >> implode)
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   908
        --| junk))
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   909
      -- Scan.repeat ((Scan.many1 is_valid >> implode) --| junk) >> op ::);
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   910
  in
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   911
    explode #> scan_valids #> implode
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   912
  end;
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   913
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   914
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   915
fun init_instantiation arities thy =
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   916
  let
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   917
    val _ = if null arities then error "At least one arity must be given" else ();
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   918
    val _ = case (duplicates (op =) o map #1) arities
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   919
     of [] => ()
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   920
      | dupl_tycos => error ("Type constructors occur more than once in arities: "
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   921
          ^ commas_quote dupl_tycos);
25502
9200b36280c0 instance command as rudimentary class target
haftmann
parents: 25485
diff changeset
   922
    val _ = map (map (the_class_data thy) o #3) arities;
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   923
    val ty_insts = map (fn (tyco, sorts, _) =>
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   924
        (tyco, Type (tyco, map TFree (Name.names Name.context Name.aT sorts))))
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   925
      arities;
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   926
    val ty_inst = the o AList.lookup (op =) ty_insts;
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   927
    fun type_name "*" = "prod"
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   928
      | type_name "+" = "sum"
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   929
      | type_name s = sanatize_name (NameSpace.base s); (*FIXME*)
25502
9200b36280c0 instance command as rudimentary class target
haftmann
parents: 25485
diff changeset
   930
    fun get_param tyco sorts (param, (c, ty)) = if can (inst thy) (c, tyco)
9200b36280c0 instance command as rudimentary class target
haftmann
parents: 25485
diff changeset
   931
      then NONE else SOME ((unoverload_const thy (c, ty), tyco),
9200b36280c0 instance command as rudimentary class target
haftmann
parents: 25485
diff changeset
   932
        (param ^ "_" ^ type_name tyco, map_atyps (K (ty_inst tyco)) ty));
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   933
    fun get_params (tyco, sorts, sort) =
25502
9200b36280c0 instance command as rudimentary class target
haftmann
parents: 25485
diff changeset
   934
      map_filter (get_param tyco sorts) (these_params thy sort)
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   935
    val params = maps get_params arities;
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   936
  in
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   937
    thy
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   938
    |> ProofContext.init
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   939
    |> Instantiation.put (mk_instantiation (arities, params))
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   940
    |> fold (Variable.declare_term o Logic.mk_type o snd) ty_insts
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   941
    |> fold ProofContext.add_arity arities
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   942
    |> Context.proof_map (
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   943
        Syntax.add_term_check 0 "instance" inst_term_check
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   944
        #> Syntax.add_term_uncheck 0 "instance" inst_term_uncheck)
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   945
  end;
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   946
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   947
fun gen_instantiation_instance do_proof after_qed lthy =
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   948
  let
25514
4b508bb31a6c first working version of instance target
haftmann
parents: 25502
diff changeset
   949
    val arities = (#arities o the_instantiation) lthy;
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   950
    val arities_proof = maps Logic.mk_arities arities;
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   951
    fun after_qed' results =
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   952
      LocalTheory.theory (fold (AxClass.add_arity o Thm.varifyT) results)
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   953
      #> after_qed;
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   954
  in
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   955
    lthy
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   956
    |> do_proof after_qed' arities_proof
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   957
  end;
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   958
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   959
val instantiation_instance = gen_instantiation_instance (fn after_qed => fn ts =>
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   960
  Proof.theorem_i NONE (after_qed o map the_single) (map (fn t => [(t, [])]) ts));
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   961
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   962
fun prove_instantiation_instance tac = gen_instantiation_instance (fn after_qed =>
25502
9200b36280c0 instance command as rudimentary class target
haftmann
parents: 25485
diff changeset
   963
  fn ts => fn lthy => after_qed (map (fn t => Goal.prove lthy [] [] t
9200b36280c0 instance command as rudimentary class target
haftmann
parents: 25485
diff changeset
   964
    (fn {context, ...} => tac context)) ts) lthy) I;
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   965
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   966
fun conclude_instantiation lthy =
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   967
  let
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   968
    val { arities, params } = the_instantiation lthy;
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   969
    val thy = ProofContext.theory_of lthy;
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   970
    (*val _ = map (fn (tyco, sorts, sort) =>
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   971
      if Sign.of_sort thy
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   972
        (Type (tyco, map TFree (Name.names Name.context Name.aT sorts)), sort)
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   973
      then () else error ("Missing instance proof for type " ^ quote (Sign.extern_type thy tyco)))
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   974
        arities; FIXME activate when old instance command is gone*)
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   975
    val params_of = maps (these o try (#params o AxClass.get_info thy))
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   976
      o Sign.complete_sort thy;
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   977
    val missing_params = arities
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   978
      |> maps (fn (tyco, _, sort) => params_of sort |> map (rpair tyco))
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   979
      |> filter_out (can (inst thy) o apfst fst);
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   980
    fun declare_missing ((c, ty0), tyco) thy =
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   981
    (*fun declare_missing ((c, tyco), (_, ty)) thy =*)
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   982
      let
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   983
        val SOME class = AxClass.class_of_param thy c;
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   984
        val name_inst = AxClass.instance_name (tyco, class) ^ "_inst";
25502
9200b36280c0 instance command as rudimentary class target
haftmann
parents: 25485
diff changeset
   985
        val c' = NameSpace.base c ^ "_" ^ NameSpace.base tyco;
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   986
        val vs = Name.names Name.context Name.aT (replicate (Sign.arity_number thy tyco) []);
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   987
        val ty = map_atyps (fn _ => Type (tyco, map TFree vs)) ty0;
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   988
      in
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   989
        thy
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   990
        |> Sign.sticky_prefix name_inst
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   991
        |> Sign.no_base_names
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   992
        |> Sign.declare_const [] (c', ty, NoSyn)
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   993
        |-> (fn const' as Const (c'', _) => Thm.add_def true
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   994
              (Thm.def_name c', Logic.mk_equals (const', Const (c, ty)))
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   995
        #>> Thm.varifyT
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   996
        #-> (fn thm => add_inst (c, tyco) (c'', Thm.symmetric thm)
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   997
        #> primitive_note Thm.internalK (c', thm)
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   998
        #> snd
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   999
        #> Sign.restore_naming thy))
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
  1000
      end;
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
  1001
  in
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
  1002
    lthy
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
  1003
    |> LocalTheory.theory (fold declare_missing missing_params)
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
  1004
  end;
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
  1005
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
  1006
end;