src/Pure/Isar/class.ML
author haftmann
Fri, 02 Jan 2009 08:13:12 +0100
changeset 29333 496b94152b55
parent 29133 9d10cc6aaa02
child 29358 efdfe5dfe008
permissions -rw-r--r--
improved boostrap order: theory_target.ML before expression.ML
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*)
26247
b6608fbeaae1 some theorems named explicitly
haftmann
parents: 26238
diff changeset
    11
  val class: bstring -> class list -> Element.context_i list
26518
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
    12
    -> theory -> string * Proof.context
26247
b6608fbeaae1 some theorems named explicitly
haftmann
parents: 26238
diff changeset
    13
  val class_cmd: bstring -> xstring list -> Element.context list
26518
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
    14
    -> 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
28017
4919bd124a58 type Properties.T;
wenzelm
parents: 27761
diff changeset
    17
  val declare: class -> Properties.T
27690
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
    18
    -> (string * mixfix) * term -> theory -> theory
28017
4919bd124a58 type Properties.T;
wenzelm
parents: 27761
diff changeset
    19
  val abbrev: class -> Syntax.mode -> Properties.T
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
26470
e44d24620515 unfold_locales now part of default tactic
haftmann
parents: 26463
diff changeset
    24
  val default_intro_tac: Proof.context -> thm list -> tactic
27684
f45fd1159a4b subclass now also works for subclasses with empty specificaton
haftmann
parents: 27281
diff changeset
    25
  val prove_subclass: class * class -> thm option -> theory -> theory
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
    26
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
    27
  val class_prefix: string -> string
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
    28
  val is_class: theory -> class -> bool
26518
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
    29
  val these_params: theory -> sort -> (string * (class * (string * typ))) list
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
    30
  val print_classes: theory -> unit
24423
ae9cd0e92423 overloaded definitions accompanied by explicit constants
haftmann
parents: 24304
diff changeset
    31
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
    32
  (*instances*)
26247
b6608fbeaae1 some theorems named explicitly
haftmann
parents: 26238
diff changeset
    33
  val init_instantiation: string list * (string * sort) list * sort
b6608fbeaae1 some theorems named explicitly
haftmann
parents: 26238
diff changeset
    34
    -> theory -> local_theory
b6608fbeaae1 some theorems named explicitly
haftmann
parents: 26238
diff changeset
    35
  val instantiation_instance: (local_theory -> local_theory)
b6608fbeaae1 some theorems named explicitly
haftmann
parents: 26238
diff changeset
    36
    -> local_theory -> Proof.state
b6608fbeaae1 some theorems named explicitly
haftmann
parents: 26238
diff changeset
    37
  val prove_instantiation_instance: (Proof.context -> tactic)
b6608fbeaae1 some theorems named explicitly
haftmann
parents: 26238
diff changeset
    38
    -> local_theory -> local_theory
28666
d2dbfe3a0284 prove_instantiation_exit combinators
haftmann
parents: 28259
diff changeset
    39
  val prove_instantiation_exit: (Proof.context -> tactic)
d2dbfe3a0284 prove_instantiation_exit combinators
haftmann
parents: 28259
diff changeset
    40
    -> local_theory -> theory
d2dbfe3a0284 prove_instantiation_exit combinators
haftmann
parents: 28259
diff changeset
    41
  val prove_instantiation_exit_result: (morphism -> 'a -> 'b)
d2dbfe3a0284 prove_instantiation_exit combinators
haftmann
parents: 28259
diff changeset
    42
    -> (Proof.context -> 'b -> tactic) -> 'a -> local_theory -> 'b * theory
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
    43
  val conclude_instantiation: local_theory -> local_theory
25603
4b7a58fc168c dropped Class.prep_spec
haftmann
parents: 25597
diff changeset
    44
  val instantiation_param: local_theory -> string -> string option
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
    45
  val confirm_declaration: string -> local_theory -> local_theory
25603
4b7a58fc168c dropped Class.prep_spec
haftmann
parents: 25597
diff changeset
    46
  val pretty_instantiation: local_theory -> Pretty.T
26259
d30f4a509361 better improvement in instantiation target
haftmann
parents: 26247
diff changeset
    47
  val type_name: string -> string
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
    48
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
    49
  (*old axclass layer*)
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
    50
  val axclass_cmd: bstring * xstring list
28084
a05ca48ef263 type Attrib.binding abbreviates Name.binding without attributes;
wenzelm
parents: 28083
diff changeset
    51
    -> (Attrib.binding * string list) list
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
    52
    -> theory -> class * theory
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
    53
  val classrel_cmd: xstring * xstring -> theory -> Proof.state
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
    54
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
    55
  (*old instance layer*)
25536
01753a944433 improved
haftmann
parents: 25518
diff changeset
    56
  val instance_arity: (theory -> theory) -> arity -> theory -> Proof.state
01753a944433 improved
haftmann
parents: 25518
diff changeset
    57
  val instance_arity_cmd: bstring * xstring list * xstring -> theory -> Proof.state
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    58
end;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    59
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    60
structure Class : CLASS =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    61
struct
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
    62
29133
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
    63
(*temporary adaption code to mediate between old and new locale code*)
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
    64
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
    65
structure Old_Locale =
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
    66
struct
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
    67
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
    68
val intro_locales_tac = Locale.intro_locales_tac; (*already forked!*)
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
    69
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
    70
val interpretation = Locale.interpretation;
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
    71
val interpretation_in_locale = Locale.interpretation_in_locale;
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
    72
val get_interpret_morph = Locale.get_interpret_morph;
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
    73
val Locale = Locale.Locale;
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
    74
val extern = Locale.extern;
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
    75
val intros = Locale.intros;
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
    76
val dests = Locale.dests;
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
    77
val init = Locale.init;
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
    78
val Merge = Locale.Merge;
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
    79
val parameters_of_expr = Locale.parameters_of_expr;
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
    80
val empty = Locale.empty;
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
    81
val cert_expr = Locale.cert_expr;
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
    82
val read_expr = Locale.read_expr;
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
    83
val parameters_of = Locale.parameters_of;
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
    84
val add_locale = Locale.add_locale;
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
    85
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
    86
end;
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
    87
29333
496b94152b55 improved boostrap order: theory_target.ML before expression.ML
haftmann
parents: 29133
diff changeset
    88
(*structure New_Locale =
29133
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
    89
struct
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
    90
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
    91
val intro_locales_tac = Locale.intro_locales_tac; (*already forked!*)
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
    92
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
    93
val interpretation = Locale.interpretation; (*!*)
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
    94
val interpretation_in_locale = Locale.interpretation_in_locale; (*!*)
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
    95
val get_interpret_morph = Locale.get_interpret_morph; (*!*)
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
    96
fun Locale loc = ([(loc, ("", Expression.Positional []))], []);
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
    97
val extern = NewLocale.extern;
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
    98
val intros = Locale.intros; (*!*)
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
    99
val dests = Locale.dests; (*!*)
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
   100
val init = NewLocale.init;
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
   101
fun Merge locs = (map (fn loc => (loc, ("", Expression.Positional []))) locs, []);
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
   102
val parameters_of_expr = Locale.parameters_of_expr; (*!*)
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
   103
val empty = ([], []);
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
   104
val cert_expr = Locale.cert_expr; (*!"*)
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
   105
val read_expr = Locale.read_expr; (*!"*)
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
   106
val parameters_of = NewLocale.params_of; (*why typ option?*)
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
   107
val add_locale = Expression.add_locale;
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
   108
29333
496b94152b55 improved boostrap order: theory_target.ML before expression.ML
haftmann
parents: 29133
diff changeset
   109
end;*)
29133
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
   110
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
   111
structure Locale = Old_Locale;
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
   112
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
   113
(*proper code again*)
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
   114
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
   115
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   116
(** auxiliary **)
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   117
25002
haftmann
parents: 24981
diff changeset
   118
fun prove_interpretation tac prfx_atts expr inst =
28822
7ca11ecbc4fb explicit name morphism function for locale interpretation
haftmann
parents: 28739
diff changeset
   119
  Locale.interpretation I prfx_atts expr inst
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   120
  ##> Proof.global_terminal_proof
27761
b95e9ba0ca1d Interpretation command (theory/proof context) no longer simplifies goal.
ballarin
parents: 27708
diff changeset
   121
      (Method.Basic (fn ctxt => Method.SIMPLE_METHOD (tac ctxt), Position.none), NONE)
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   122
  ##> ProofContext.theory_of;
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   123
25195
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   124
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
   125
  Locale.interpretation_in_locale
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   126
      (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
   127
  #> Proof.global_terminal_proof
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   128
      (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
   129
  #> ProofContext.theory_of;
62638dcafe38 fixed syntax; truned code structure; added primitive subclass interface with consideraton of syntax etc.
haftmann
parents: 25163
diff changeset
   130
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   131
val class_prefix = Logic.const_of_class o Sign.base_name;
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   132
28822
7ca11ecbc4fb explicit name morphism function for locale interpretation
haftmann
parents: 28739
diff changeset
   133
fun class_name_morphism class =
28941
128459bd72d2 new Binding module
haftmann
parents: 28861
diff changeset
   134
  Binding.map_prefix (K (Binding.add_prefix false (class_prefix class)));
28822
7ca11ecbc4fb explicit name morphism function for locale interpretation
haftmann
parents: 28739
diff changeset
   135
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   136
fun activate_class_morphism thy class inst morphism =
28822
7ca11ecbc4fb explicit name morphism function for locale interpretation
haftmann
parents: 28739
diff changeset
   137
  Locale.get_interpret_morph thy (class_name_morphism class) (class, "") morphism class inst;
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   138
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   139
fun prove_class_interpretation class inst consts hyp_facts def_facts thy =
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   140
  let
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   141
    val constraints = map (fn (class, c) => map_atyps (K (TFree (Name.aT,
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   142
      [class]))) (Sign.the_const_type thy c)) consts;
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   143
    val no_constraints = map (map_atyps (K (TFree (Name.aT, [])))) constraints;
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   144
    fun add_constraint c T = Sign.add_const_constraint (c, SOME T);
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   145
    fun tac ctxt = ALLGOALS (ProofContext.fact_tac (hyp_facts @ def_facts)
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   146
      ORELSE' (fn n => SELECT_GOAL (Locale.intro_locales_tac false ctxt []) n));
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   147
    val prfx = class_prefix class;
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   148
  in
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   149
    thy
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   150
    |> fold2 add_constraint (map snd consts) no_constraints
28822
7ca11ecbc4fb explicit name morphism function for locale interpretation
haftmann
parents: 28739
diff changeset
   151
    |> prove_interpretation tac (class_name_morphism class) (Locale.Locale class)
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28941
diff changeset
   152
          (map SOME inst, map (pair (Attrib.empty_binding) o Thm.prop_of) def_facts)
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   153
    ||> fold2 add_constraint (map snd consts) constraints
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   154
  end;
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   155
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   156
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   157
(** primitive axclass and instance commands **)
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   158
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   159
fun axclass_cmd (class, raw_superclasses) raw_specs thy =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   160
  let
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   161
    val ctxt = ProofContext.init thy;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   162
    val superclasses = map (Sign.read_class thy) raw_superclasses;
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   163
    val name_atts = map ((apsnd o map) (Attrib.attribute thy) o fst)
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   164
      raw_specs;
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   165
    val axiomss = ProofContext.read_propp (ctxt, map (map (rpair []) o snd)
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   166
          raw_specs)
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   167
      |> snd
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   168
      |> (map o map) fst;
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   169
  in
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   170
    AxClass.define_class (class, superclasses) []
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   171
      (name_atts ~~ axiomss) thy
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   172
  end;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   173
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   174
local
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   175
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   176
fun gen_instance mk_prop add_thm after_qed insts thy =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   177
  let
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   178
    fun after_qed' results =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   179
      ProofContext.theory ((fold o fold) add_thm results #> after_qed);
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   180
  in
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   181
    thy
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   182
    |> ProofContext.init
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   183
    |> Proof.theorem_i NONE after_qed' ((map (fn t => [(t, [])])
25536
01753a944433 improved
haftmann
parents: 25518
diff changeset
   184
        o mk_prop thy) insts)
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   185
  end;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   186
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   187
in
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   188
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   189
val instance_arity =
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   190
  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
   191
val instance_arity_cmd =
9200b36280c0 instance command as rudimentary class target
haftmann
parents: 25485
diff changeset
   192
  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
   193
val classrel =
25536
01753a944433 improved
haftmann
parents: 25518
diff changeset
   194
  gen_instance (single oo (Logic.mk_classrel oo AxClass.cert_classrel)) AxClass.add_classrel I;
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   195
val classrel_cmd =
25536
01753a944433 improved
haftmann
parents: 25518
diff changeset
   196
  gen_instance (single oo (Logic.mk_classrel oo AxClass.read_classrel)) AxClass.add_classrel I;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   197
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   198
end; (*local*)
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   199
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   200
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   201
(** class data **)
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   202
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   203
datatype class_data = ClassData of {
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   204
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   205
  (* static part *)
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   206
  consts: (string * string) list
24836
dab06e93ec28 intermediate cleanup
haftmann
parents: 24770
diff changeset
   207
    (*locale parameter ~> constant name*),
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   208
  base_sort: sort,
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   209
  inst: term list
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   210
    (*canonical interpretation*),
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   211
  morphism: Morphism.morphism * ((typ Vartab.table * typ list) * (term Vartab.table * term list))
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   212
    (*morphism cookie of canonical interpretation*),
25618
01f20279fea1 improved rule calculation
haftmann
parents: 25603
diff changeset
   213
  assm_intro: thm option,
01f20279fea1 improved rule calculation
haftmann
parents: 25603
diff changeset
   214
  of_class: thm,
01f20279fea1 improved rule calculation
haftmann
parents: 25603
diff changeset
   215
  axiom: thm option,
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   216
  
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   217
  (* dynamic part *)
24657
185502d54c3d fixed wrong syntax treatment in class target
haftmann
parents: 24589
diff changeset
   218
  defs: thm list,
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   219
  operations: (string * (class * (typ * term))) list
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   220
24657
185502d54c3d fixed wrong syntax treatment in class target
haftmann
parents: 24589
diff changeset
   221
};
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   222
24657
185502d54c3d fixed wrong syntax treatment in class target
haftmann
parents: 24589
diff changeset
   223
fun rep_class_data (ClassData d) = d;
25618
01f20279fea1 improved rule calculation
haftmann
parents: 25603
diff changeset
   224
fun mk_class_data ((consts, base_sort, inst, morphism, assm_intro, of_class, axiom),
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   225
    (defs, operations)) =
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   226
  ClassData { consts = consts, base_sort = base_sort, inst = inst,
26463
9283b4185fdf Context.>> : operate on Context.generic;
wenzelm
parents: 26435
diff changeset
   227
    morphism = morphism, assm_intro = assm_intro, of_class = of_class, axiom = axiom,
25618
01f20279fea1 improved rule calculation
haftmann
parents: 25603
diff changeset
   228
    defs = defs, operations = operations };
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   229
fun map_class_data f (ClassData { consts, base_sort, inst, morphism, assm_intro,
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   230
    of_class, axiom, defs, operations }) =
25618
01f20279fea1 improved rule calculation
haftmann
parents: 25603
diff changeset
   231
  mk_class_data (f ((consts, base_sort, inst, morphism, assm_intro, of_class, axiom),
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   232
    (defs, operations)));
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   233
fun merge_class_data _ (ClassData { consts = consts,
25618
01f20279fea1 improved rule calculation
haftmann
parents: 25603
diff changeset
   234
    base_sort = base_sort, inst = inst, morphism = morphism, assm_intro = assm_intro,
01f20279fea1 improved rule calculation
haftmann
parents: 25603
diff changeset
   235
    of_class = of_class, axiom = axiom, defs = defs1, operations = operations1 },
01f20279fea1 improved rule calculation
haftmann
parents: 25603
diff changeset
   236
  ClassData { consts = _, base_sort = _, inst = _, morphism = _, assm_intro = _,
01f20279fea1 improved rule calculation
haftmann
parents: 25603
diff changeset
   237
    of_class = _, axiom = _, defs = defs2, operations = operations2 }) =
01f20279fea1 improved rule calculation
haftmann
parents: 25603
diff changeset
   238
  mk_class_data ((consts, base_sort, inst, morphism, assm_intro, of_class, axiom),
24914
95cda5dd58d5 added proper subclass concept; improved class target
haftmann
parents: 24901
diff changeset
   239
    (Thm.merge_thms (defs1, defs2),
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   240
      AList.merge (op =) (K true) (operations1, operations2)));
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   241
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   242
structure ClassData = TheoryDataFun
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   243
(
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   244
  type T = class_data Graph.T
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   245
  val empty = Graph.empty;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   246
  val copy = I;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   247
  val extend = I;
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   248
  fun merge _ = Graph.join merge_class_data;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   249
);
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   250
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   251
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   252
(* queries *)
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   253
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   254
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
   255
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   256
fun the_class_data thy class = case lookup_class_data thy class
25020
f1344290e420 (un)overload: full rewrite;
wenzelm
parents: 25010
diff changeset
   257
 of NONE => error ("Undeclared class " ^ quote class)
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   258
  | SOME data => data;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   259
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   260
val is_class = is_some oo lookup_class_data;
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   261
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   262
val ancestry = Graph.all_succs o ClassData.get;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   263
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   264
fun the_inst thy = #inst o the_class_data thy;
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   265
25002
haftmann
parents: 24981
diff changeset
   266
fun these_params thy =
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   267
  let
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   268
    fun params class =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   269
      let
24930
cc2e0e8c81af renamed AxClass.get_definition to AxClass.get_info (again);
wenzelm
parents: 24920
diff changeset
   270
        val const_typs = (#params o AxClass.get_info thy) class;
24657
185502d54c3d fixed wrong syntax treatment in class target
haftmann
parents: 24589
diff changeset
   271
        val const_names = (#consts o the_class_data thy) class;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   272
      in
26518
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   273
        (map o apsnd)
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   274
          (fn c => (class, (c, (the o AList.lookup (op =) const_typs) c))) const_names
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   275
      end;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   276
  in maps params o ancestry thy end;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   277
25618
01f20279fea1 improved rule calculation
haftmann
parents: 25603
diff changeset
   278
fun these_assm_intros thy =
01f20279fea1 improved rule calculation
haftmann
parents: 25603
diff changeset
   279
  Graph.fold (fn (_, (data, _)) => fold (insert Thm.eq_thm)
01f20279fea1 improved rule calculation
haftmann
parents: 25603
diff changeset
   280
    ((the_list o #assm_intro o rep_class_data) data)) (ClassData.get thy) [];
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   281
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   282
fun these_defs thy = maps (these o Option.map #defs o lookup_class_data thy) o ancestry thy;
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   283
24836
dab06e93ec28 intermediate cleanup
haftmann
parents: 24770
diff changeset
   284
fun these_operations thy =
dab06e93ec28 intermediate cleanup
haftmann
parents: 24770
diff changeset
   285
  maps (#operations o the_class_data thy) o ancestry thy;
24657
185502d54c3d fixed wrong syntax treatment in class target
haftmann
parents: 24589
diff changeset
   286
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   287
fun morphism thy class =
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   288
  let
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   289
    val { inst, morphism, ... } = the_class_data thy class;
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   290
  in activate_class_morphism thy class inst morphism end;
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   291
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   292
fun print_classes thy =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   293
  let
24920
2a45e400fdad generic Syntax.pretty/string_of operations;
wenzelm
parents: 24914
diff changeset
   294
    val ctxt = ProofContext.init thy;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   295
    val algebra = Sign.classes_of thy;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   296
    val arities =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   297
      Symtab.empty
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   298
      |> Symtab.fold (fn (tyco, arities) => fold (fn (class, _) =>
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   299
           Symtab.map_default (class, []) (insert (op =) tyco)) arities)
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   300
             ((#arities o Sorts.rep_algebra) algebra);
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   301
    val the_arities = these o Symtab.lookup arities;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   302
    fun mk_arity class tyco =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   303
      let
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   304
        val Ss = Sorts.mg_domain algebra tyco [class];
24920
2a45e400fdad generic Syntax.pretty/string_of operations;
wenzelm
parents: 24914
diff changeset
   305
      in Syntax.pretty_arity ctxt (tyco, Ss, [class]) end;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   306
    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
   307
      ^ setmp show_sorts false (Syntax.string_of_typ ctxt o Type.strip_sorts) ty);
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   308
    fun mk_entry class = (Pretty.block o Pretty.fbreaks o map_filter I) [
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   309
      (SOME o Pretty.str) ("class " ^ Sign.extern_class thy class ^ ":"),
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   310
      (SOME o Pretty.block) [Pretty.str "supersort: ",
24920
2a45e400fdad generic Syntax.pretty/string_of operations;
wenzelm
parents: 24914
diff changeset
   311
        (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
   312
      if is_class thy class then (SOME o Pretty.str)
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   313
        ("locale: " ^ Locale.extern thy class) else NONE,
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   314
      ((fn [] => NONE | ps => (SOME o Pretty.block o Pretty.fbreaks)
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   315
          (Pretty.str "parameters:" :: ps)) o map mk_param
24930
cc2e0e8c81af renamed AxClass.get_definition to AxClass.get_info (again);
wenzelm
parents: 24920
diff changeset
   316
        o these o Option.map #params o try (AxClass.get_info thy)) class,
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   317
      (SOME o Pretty.block o Pretty.breaks) [
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   318
        Pretty.str "instances:",
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   319
        Pretty.list "" "" (map (mk_arity class) (the_arities class))
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   320
      ]
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   321
    ]
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   322
  in
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   323
    (Pretty.writeln o Pretty.chunks o separate (Pretty.str "")
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   324
      o map mk_entry o Sorts.all_classes) algebra
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   325
  end;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   326
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   327
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   328
(* updaters *)
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   329
25618
01f20279fea1 improved rule calculation
haftmann
parents: 25603
diff changeset
   330
fun add_class_data ((class, superclasses),
25711
91cee0cefaf7 tuned primitive inferences
haftmann
parents: 25683
diff changeset
   331
    (params, base_sort, inst, phi, axiom, assm_intro, of_class)) thy =
25002
haftmann
parents: 24981
diff changeset
   332
  let
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   333
    val operations = map (fn (v_ty as (_, ty), (c, _)) =>
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   334
      (c, (class, (ty, Free v_ty)))) params;
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   335
    val add_class = Graph.new_node (class,
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   336
        mk_class_data (((map o pairself) fst params, base_sort,
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   337
          inst, phi, assm_intro, of_class, axiom), ([], operations)))
25002
haftmann
parents: 24981
diff changeset
   338
      #> fold (curry Graph.add_edge class) superclasses;
25618
01f20279fea1 improved rule calculation
haftmann
parents: 25603
diff changeset
   339
  in ClassData.map add_class thy end;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   340
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   341
fun register_operation class (c, (t, some_def)) thy =
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   342
  let
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   343
    val base_sort = (#base_sort o the_class_data thy) class;
26518
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   344
    val prep_typ = map_type_tvar
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   345
      (fn (vi as (v, _), sort) => if Name.aT = v
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   346
        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
   347
    val t' = map_types prep_typ t;
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   348
    val ty' = Term.fastype_of t';
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   349
  in
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   350
    thy
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   351
    |> (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
   352
      (fn (defs, operations) =>
25096
b8950f7cf92e clarified abbreviations in class context
haftmann
parents: 25094
diff changeset
   353
        (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
   354
          (c, (class, (ty', t'))) :: operations))
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   355
  end;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   356
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   357
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   358
(** rule calculation, tactics and methods **)
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   359
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   360
fun calculate_axiom thy sups base_sort assm_axiom param_map class =
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   361
  case Locale.intros thy class
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   362
   of (_, []) => assm_axiom
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   363
    | (_, [intro]) =>
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   364
      let
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   365
        fun instantiate thy sort = Thm.instantiate ([pairself (Thm.ctyp_of thy o TVar o pair (Name.aT, 0))
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   366
          (base_sort, sort)], map (fn (v, (c, ty)) => pairself (Thm.cterm_of thy)
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   367
            (Var ((v, 0), map_atyps (fn _ => TVar ((Name.aT, 0), sort)) ty),
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   368
              Const (c, map_atyps (fn _ => TVar ((Name.aT, 0), sort)) ty))) param_map);
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   369
        val axiom_premises = map_filter (#axiom o the_class_data thy) sups
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   370
          @ the_list assm_axiom;
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   371
      in intro
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   372
        |> instantiate thy [class]
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   373
        |> (fn thm => thm OF axiom_premises)
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   374
        |> Drule.standard'
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   375
        |> Thm.close_derivation
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   376
        |> SOME
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   377
      end;
25024
0615bb9955dd added is_class;
wenzelm
parents: 25020
diff changeset
   378
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   379
fun calculate_rules thy phi sups base_sort param_map axiom class =
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   380
  let
25711
91cee0cefaf7 tuned primitive inferences
haftmann
parents: 25683
diff changeset
   381
    fun instantiate thy sort = Thm.instantiate ([pairself (Thm.ctyp_of thy o TVar o pair (Name.aT, 0))
91cee0cefaf7 tuned primitive inferences
haftmann
parents: 25683
diff changeset
   382
      (base_sort, sort)], map (fn (v, (c, ty)) => pairself (Thm.cterm_of thy)
91cee0cefaf7 tuned primitive inferences
haftmann
parents: 25683
diff changeset
   383
        (Var ((v, 0), map_atyps (fn _ => TVar ((Name.aT, 0), sort)) ty),
91cee0cefaf7 tuned primitive inferences
haftmann
parents: 25683
diff changeset
   384
          Const (c, map_atyps (fn _ => TVar ((Name.aT, 0), sort)) ty))) param_map);
91cee0cefaf7 tuned primitive inferences
haftmann
parents: 25683
diff changeset
   385
    val defs = these_defs thy sups;
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   386
    val assm_intro = Locale.intros thy class
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   387
      |> fst
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   388
      |> map (instantiate thy base_sort)
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   389
      |> map (MetaSimplifier.rewrite_rule defs)
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   390
      |> map Thm.close_derivation
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   391
      |> try the_single;
25711
91cee0cefaf7 tuned primitive inferences
haftmann
parents: 25683
diff changeset
   392
    val fixate = Thm.instantiate
91cee0cefaf7 tuned primitive inferences
haftmann
parents: 25683
diff changeset
   393
      (map (pairself (Thm.ctyp_of thy)) [(TVar ((Name.aT, 0), []), TFree (Name.aT, base_sort)),
91cee0cefaf7 tuned primitive inferences
haftmann
parents: 25683
diff changeset
   394
        (TVar ((Name.aT, 0), base_sort), TFree (Name.aT, base_sort))], [])
25618
01f20279fea1 improved rule calculation
haftmann
parents: 25603
diff changeset
   395
    val of_class_sups = if null sups
25711
91cee0cefaf7 tuned primitive inferences
haftmann
parents: 25683
diff changeset
   396
      then map (fixate o Thm.class_triv thy) base_sort
91cee0cefaf7 tuned primitive inferences
haftmann
parents: 25683
diff changeset
   397
      else map (fixate o #of_class o the_class_data thy) sups;
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   398
    val locale_dests = map Drule.standard' (Locale.dests thy class);
25711
91cee0cefaf7 tuned primitive inferences
haftmann
parents: 25683
diff changeset
   399
    val num_trivs = case length locale_dests
91cee0cefaf7 tuned primitive inferences
haftmann
parents: 25683
diff changeset
   400
     of 0 => if is_none axiom then 0 else 1
91cee0cefaf7 tuned primitive inferences
haftmann
parents: 25683
diff changeset
   401
      | n => n;
91cee0cefaf7 tuned primitive inferences
haftmann
parents: 25683
diff changeset
   402
    val pred_trivs = if num_trivs = 0 then []
91cee0cefaf7 tuned primitive inferences
haftmann
parents: 25683
diff changeset
   403
      else the axiom
91cee0cefaf7 tuned primitive inferences
haftmann
parents: 25683
diff changeset
   404
        |> Thm.prop_of
91cee0cefaf7 tuned primitive inferences
haftmann
parents: 25683
diff changeset
   405
        |> (map_types o map_atyps o K) (TFree (Name.aT, base_sort))
91cee0cefaf7 tuned primitive inferences
haftmann
parents: 25683
diff changeset
   406
        |> (Thm.assume o Thm.cterm_of thy)
91cee0cefaf7 tuned primitive inferences
haftmann
parents: 25683
diff changeset
   407
        |> replicate num_trivs;
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   408
    val axclass_intro = (#intro o AxClass.get_info thy) class;
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   409
    val of_class = (fixate axclass_intro OF of_class_sups OF locale_dests OF pred_trivs)
25711
91cee0cefaf7 tuned primitive inferences
haftmann
parents: 25683
diff changeset
   410
      |> Drule.standard'
26628
63306cb94313 replaced Drule.close_derivation/Goal.close_result by Thm.close_derivation (removed obsolete compression);
wenzelm
parents: 26596
diff changeset
   411
      |> Thm.close_derivation;
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   412
  in (assm_intro, of_class) end;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   413
27684
f45fd1159a4b subclass now also works for subclasses with empty specificaton
haftmann
parents: 27281
diff changeset
   414
fun prove_subclass (sub, sup) some_thm thy =
25618
01f20279fea1 improved rule calculation
haftmann
parents: 25603
diff changeset
   415
  let
25711
91cee0cefaf7 tuned primitive inferences
haftmann
parents: 25683
diff changeset
   416
    val of_class = (#of_class o the_class_data thy) sup;
27684
f45fd1159a4b subclass now also works for subclasses with empty specificaton
haftmann
parents: 27281
diff changeset
   417
    val intro = case some_thm
f45fd1159a4b subclass now also works for subclasses with empty specificaton
haftmann
parents: 27281
diff changeset
   418
     of SOME thm => Drule.standard' (of_class OF [Drule.standard' thm])
f45fd1159a4b subclass now also works for subclasses with empty specificaton
haftmann
parents: 27281
diff changeset
   419
      | NONE => Thm.instantiate ([pairself (Thm.ctyp_of thy o TVar o pair (Name.aT, 0))
f45fd1159a4b subclass now also works for subclasses with empty specificaton
haftmann
parents: 27281
diff changeset
   420
          ([], [sub])], []) of_class;
f45fd1159a4b subclass now also works for subclasses with empty specificaton
haftmann
parents: 27281
diff changeset
   421
    val classrel = (intro OF (the_list o #axiom o the_class_data thy) sub)
f45fd1159a4b subclass now also works for subclasses with empty specificaton
haftmann
parents: 27281
diff changeset
   422
      |> Thm.close_derivation;
25618
01f20279fea1 improved rule calculation
haftmann
parents: 25603
diff changeset
   423
  in
01f20279fea1 improved rule calculation
haftmann
parents: 25603
diff changeset
   424
    thy
01f20279fea1 improved rule calculation
haftmann
parents: 25603
diff changeset
   425
    |> AxClass.add_classrel classrel
27684
f45fd1159a4b subclass now also works for subclasses with empty specificaton
haftmann
parents: 27281
diff changeset
   426
    |> prove_interpretation_in (ALLGOALS (ProofContext.fact_tac (the_list some_thm)))
25618
01f20279fea1 improved rule calculation
haftmann
parents: 25603
diff changeset
   427
         I (sub, Locale.Locale sup)
01f20279fea1 improved rule calculation
haftmann
parents: 25603
diff changeset
   428
    |> ClassData.map (Graph.add_edge (sub, sup))
01f20279fea1 improved rule calculation
haftmann
parents: 25603
diff changeset
   429
  end;
01f20279fea1 improved rule calculation
haftmann
parents: 25603
diff changeset
   430
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   431
fun note_assm_intro class assm_intro thy =
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   432
  thy
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   433
  |> Sign.sticky_prefix (class_prefix class ^ "_" ^ AxClass.axiomsN)
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   434
  |> PureThy.store_thm (AxClass.introN, assm_intro)
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   435
  |> snd
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   436
  |> Sign.restore_naming thy;
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   437
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   438
fun intro_classes_tac facts st =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   439
  let
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   440
    val thy = Thm.theory_of_thm st;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   441
    val classes = Sign.all_classes thy;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   442
    val class_trivs = map (Thm.class_triv thy) classes;
25618
01f20279fea1 improved rule calculation
haftmann
parents: 25603
diff changeset
   443
    val class_intros = map_filter (try (#intro o AxClass.get_info thy)) classes;
01f20279fea1 improved rule calculation
haftmann
parents: 25603
diff changeset
   444
    val assm_intros = these_assm_intros thy;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   445
  in
25618
01f20279fea1 improved rule calculation
haftmann
parents: 25603
diff changeset
   446
    Method.intros_tac (class_trivs @ class_intros @ assm_intros) facts st
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   447
  end;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   448
26470
e44d24620515 unfold_locales now part of default tactic
haftmann
parents: 26463
diff changeset
   449
fun default_intro_tac ctxt [] =
29029
1945e0185097 NewLocale.intro_locales_tac added to Class.default_intro_tac.
ballarin
parents: 29006
diff changeset
   450
      intro_classes_tac [] ORELSE NewLocale.intro_locales_tac true ctxt [] ORELSE
1945e0185097 NewLocale.intro_locales_tac added to Class.default_intro_tac.
ballarin
parents: 29006
diff changeset
   451
      Locale.intro_locales_tac true ctxt []
26470
e44d24620515 unfold_locales now part of default tactic
haftmann
parents: 26463
diff changeset
   452
  | default_intro_tac _ _ = no_tac;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   453
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   454
fun default_tac rules ctxt facts =
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   455
  HEADGOAL (Method.some_rule_tac rules ctxt facts) ORELSE
26470
e44d24620515 unfold_locales now part of default tactic
haftmann
parents: 26463
diff changeset
   456
    default_intro_tac ctxt facts;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   457
26463
9283b4185fdf Context.>> : operate on Context.generic;
wenzelm
parents: 26435
diff changeset
   458
val _ = Context.>> (Context.map_theory
9283b4185fdf Context.>> : operate on Context.generic;
wenzelm
parents: 26435
diff changeset
   459
  (Method.add_methods
9283b4185fdf Context.>> : operate on Context.generic;
wenzelm
parents: 26435
diff changeset
   460
   [("intro_classes", Method.no_args (Method.METHOD intro_classes_tac),
9283b4185fdf Context.>> : operate on Context.generic;
wenzelm
parents: 26435
diff changeset
   461
      "back-chain introduction rules of classes"),
9283b4185fdf Context.>> : operate on Context.generic;
wenzelm
parents: 26435
diff changeset
   462
    ("default", Method.thms_ctxt_args (Method.METHOD oo default_tac),
9283b4185fdf Context.>> : operate on Context.generic;
wenzelm
parents: 26435
diff changeset
   463
      "apply some intro/elim rule")]));
9283b4185fdf Context.>> : operate on Context.generic;
wenzelm
parents: 26435
diff changeset
   464
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   465
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   466
(** classes and class target **)
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   467
25002
haftmann
parents: 24981
diff changeset
   468
(* class context syntax *)
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   469
26238
c30bb8182da2 generic improvable syntax for targets
haftmann
parents: 26167
diff changeset
   470
fun synchronize_class_syntax sups base_sort ctxt =
24914
95cda5dd58d5 added proper subclass concept; improved class target
haftmann
parents: 24901
diff changeset
   471
  let
25344
00c2179db769 synchronize_syntax: improved declare_const (still inactive);
wenzelm
parents: 25326
diff changeset
   472
    val thy = ProofContext.theory_of ctxt;
26596
07d7d0a6d5fd check validity of class target improvement
haftmann
parents: 26518
diff changeset
   473
    val algebra = Sign.classes_of thy;
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   474
    val operations = these_operations thy sups;
26518
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   475
    fun subst_class_typ sort = map_type_tfree (K (TVar ((Name.aT, 0), sort)));
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   476
    val primary_constraints =
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   477
      (map o apsnd) (subst_class_typ base_sort o fst o snd) operations;
26518
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   478
    val secondary_constraints =
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   479
      (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
   480
    fun declare_const (c, _) =
c8352b38d47d synchronize_syntax: declare operations within the local scope of fixes/consts
wenzelm
parents: 25311
diff changeset
   481
      let val b = Sign.base_name c
25344
00c2179db769 synchronize_syntax: improved declare_const (still inactive);
wenzelm
parents: 25326
diff changeset
   482
      in Sign.intern_const thy b = c ? Variable.declare_const (b, c) end;
26518
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   483
    fun improve (c, ty) = (case AList.lookup (op =) primary_constraints c
26238
c30bb8182da2 generic improvable syntax for targets
haftmann
parents: 26167
diff changeset
   484
     of SOME ty' => (case try (Type.raw_match (ty', ty)) Vartab.empty
c30bb8182da2 generic improvable syntax for targets
haftmann
parents: 26167
diff changeset
   485
         of SOME tyenv => (case Vartab.lookup tyenv (Name.aT, 0)
26596
07d7d0a6d5fd check validity of class target improvement
haftmann
parents: 26518
diff changeset
   486
             of SOME (_, ty' as TVar (tvar as (vi, sort))) =>
26238
c30bb8182da2 generic improvable syntax for targets
haftmann
parents: 26167
diff changeset
   487
                  if TypeInfer.is_param vi
26596
07d7d0a6d5fd check validity of class target improvement
haftmann
parents: 26518
diff changeset
   488
                    andalso Sorts.sort_le algebra (base_sort, sort)
07d7d0a6d5fd check validity of class target improvement
haftmann
parents: 26518
diff changeset
   489
                      then SOME (ty', TFree (Name.aT, base_sort))
07d7d0a6d5fd check validity of class target improvement
haftmann
parents: 26518
diff changeset
   490
                      else NONE
26238
c30bb8182da2 generic improvable syntax for targets
haftmann
parents: 26167
diff changeset
   491
              | _ => NONE)
c30bb8182da2 generic improvable syntax for targets
haftmann
parents: 26167
diff changeset
   492
          | NONE => NONE)
c30bb8182da2 generic improvable syntax for targets
haftmann
parents: 26167
diff changeset
   493
      | NONE => NONE)
c30bb8182da2 generic improvable syntax for targets
haftmann
parents: 26167
diff changeset
   494
    fun subst (c, ty) = Option.map snd (AList.lookup (op =) operations c);
25368
f12613fda79d proper implementation of check phase; non-qualified names for class operations
haftmann
parents: 25344
diff changeset
   495
    val unchecks = map (fn (c, (_, (ty, t))) => (t, Const (c, ty))) operations;
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   496
  in
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   497
    ctxt
26518
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   498
    |> fold declare_const primary_constraints
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   499
    |> Overloading.map_improvable_syntax (K (((primary_constraints, secondary_constraints),
26730
bbb5e6904d78 proper abbreviations in class
haftmann
parents: 26642
diff changeset
   500
        (((improve, subst), true), unchecks)), false))
26518
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   501
    |> Overloading.set_primary_constraints
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   502
  end;
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   503
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   504
fun refresh_syntax class ctxt =
25002
haftmann
parents: 24981
diff changeset
   505
  let
haftmann
parents: 24981
diff changeset
   506
    val thy = ProofContext.theory_of ctxt;
25062
af5ef0d4d655 global class syntax
haftmann
parents: 25060
diff changeset
   507
    val base_sort = (#base_sort o the_class_data thy) class;
26238
c30bb8182da2 generic improvable syntax for targets
haftmann
parents: 26167
diff changeset
   508
  in synchronize_class_syntax [class] base_sort ctxt end;
25002
haftmann
parents: 24981
diff changeset
   509
25344
00c2179db769 synchronize_syntax: improved declare_const (still inactive);
wenzelm
parents: 25326
diff changeset
   510
fun init_ctxt sups base_sort ctxt =
25083
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   511
  ctxt
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   512
  |> Variable.declare_term
765528b4b419 improved class syntax
haftmann
parents: 25066
diff changeset
   513
      (Logic.mk_type (TFree (Name.aT, base_sort)))
26238
c30bb8182da2 generic improvable syntax for targets
haftmann
parents: 26167
diff changeset
   514
  |> synchronize_class_syntax sups base_sort
c30bb8182da2 generic improvable syntax for targets
haftmann
parents: 26167
diff changeset
   515
  |> Overloading.add_improvable_syntax;
24901
d3cbf79769b9 added first version of user-space type system for class target
haftmann
parents: 24847
diff changeset
   516
25311
aa750e3a581c Class.init now similiar to Locale.init
haftmann
parents: 25268
diff changeset
   517
fun init class thy =
aa750e3a581c Class.init now similiar to Locale.init
haftmann
parents: 25268
diff changeset
   518
  thy
aa750e3a581c Class.init now similiar to Locale.init
haftmann
parents: 25268
diff changeset
   519
  |> Locale.init class
25344
00c2179db769 synchronize_syntax: improved declare_const (still inactive);
wenzelm
parents: 25326
diff changeset
   520
  |> 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
   521
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   522
27690
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   523
(* class target *)
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   524
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   525
fun declare class pos ((c, mx), dict) thy =
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   526
  let
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   527
    val prfx = class_prefix class;
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   528
    val thy' = thy |> Sign.add_path prfx;
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   529
    val phi = morphism thy' class;
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   530
    val inst = the_inst thy' class;
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   531
    val params = map (apsnd fst o snd) (these_params thy' [class]);
27690
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   532
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28941
diff changeset
   533
    val c' = Sign.full_bname thy' c;
27690
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   534
    val dict' = Morphism.term phi dict;
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   535
    val dict_def = map_types Logic.unvarifyT dict';
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   536
    val ty' = Term.fastype_of dict_def;
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   537
    val ty'' = Type.strip_sorts ty';
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   538
    val def_eq = Logic.mk_equals (Const (c', ty'), dict_def);
28674
08a77c495dc1 renamed Thm.get_axiom_i to Thm.axiom;
wenzelm
parents: 28666
diff changeset
   539
    fun get_axiom thy = ((Thm.varifyT o Thm.axiom thy) c', thy);
27690
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   540
  in
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   541
    thy'
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28941
diff changeset
   542
    |> Sign.declare_const pos ((Binding.name c, ty''), mx) |> snd
27690
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   543
    |> Thm.add_def false false (c, def_eq)
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   544
    |>> Thm.symmetric
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   545
    ||>> get_axiom
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   546
    |-> (fn (def, def') => prove_class_interpretation class inst params [] [def]
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   547
      #> snd
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   548
        (*assumption: interpretation cookie does not change
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   549
          by adding equations to interpretation*)
27690
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   550
      #> register_operation class (c', (dict', SOME (Thm.symmetric def')))
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   551
      #> PureThy.store_thm (c ^ "_raw", def')
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   552
      #> snd)
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   553
    |> Sign.restore_naming thy
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   554
    |> Sign.add_const_constraint (c', SOME ty')
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   555
  end;
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   556
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   557
fun abbrev class prmode pos ((c, mx), rhs) thy =
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   558
  let
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   559
    val prfx = class_prefix class;
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   560
    val thy' = thy |> Sign.add_path prfx;
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   561
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   562
    val unchecks = map (fn (c, (_, (ty, t))) => (t, Const (c, ty)))
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   563
      (these_operations thy [class]);
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28941
diff changeset
   564
    val c' = Sign.full_bname thy' c;
27690
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   565
    val rhs' = Pattern.rewrite_term thy unchecks [] rhs;
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   566
    val rhs'' = map_types Logic.varifyT rhs';
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   567
    val ty' = Term.fastype_of rhs';
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   568
  in
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   569
    thy'
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28941
diff changeset
   570
    |> Sign.add_abbrev (#1 prmode) pos (Binding.name c, map_types Type.strip_sorts rhs'') |> snd
27690
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   571
    |> Sign.add_const_constraint (c', SOME ty')
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   572
    |> Sign.notation true prmode [(Const (c', ty'), mx)]
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   573
    |> not (#1 prmode = PrintMode.input) ? register_operation class (c', (rhs', NONE))
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   574
    |> Sign.restore_naming thy
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   575
  end;
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   576
24738db98d34 some steps towards explicit class target for canonical interpretation
haftmann
parents: 27684
diff changeset
   577
24589
d3fca349736c clarified class interfaces and internals
haftmann
parents: 24435
diff changeset
   578
(* class definition *)
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   579
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   580
local
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   581
26247
b6608fbeaae1 some theorems named explicitly
haftmann
parents: 26238
diff changeset
   582
fun gen_class_spec prep_class process_expr thy raw_supclasses raw_elems =
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   583
  let
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   584
    val supclasses = map (prep_class thy) raw_supclasses;
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   585
    val supsort = Sign.minimize_sort thy supclasses;
25618
01f20279fea1 improved rule calculation
haftmann
parents: 25603
diff changeset
   586
    val sups = filter (is_class thy) supsort;
26995
2511a72dd73c check for illegal merge of class parameters
haftmann
parents: 26939
diff changeset
   587
    val supparam_names = map fst (these_params thy sups);
2511a72dd73c check for illegal merge of class parameters
haftmann
parents: 26939
diff changeset
   588
    val _ = if has_duplicates (op =) supparam_names
2511a72dd73c check for illegal merge of class parameters
haftmann
parents: 26939
diff changeset
   589
      then error ("Duplicate parameter(s) in superclasses: "
2511a72dd73c check for illegal merge of class parameters
haftmann
parents: 26939
diff changeset
   590
        ^ (commas o map quote o duplicates (op =)) supparam_names)
2511a72dd73c check for illegal merge of class parameters
haftmann
parents: 26939
diff changeset
   591
      else ();
25618
01f20279fea1 improved rule calculation
haftmann
parents: 25603
diff changeset
   592
    val base_sort = if null sups then supsort else
26167
ccc9007a7164 proper merge of base sorts
haftmann
parents: 25999
diff changeset
   593
      foldr1 (Sorts.inter_sort (Sign.classes_of thy))
ccc9007a7164 proper merge of base sorts
haftmann
parents: 25999
diff changeset
   594
        (map (#base_sort o the_class_data thy) sups);
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   595
    val suplocales = map Locale.Locale sups;
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   596
    val supexpr = Locale.Merge suplocales;
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   597
    val supparams = (map fst o Locale.parameters_of_expr thy) supexpr;
29133
9d10cc6aaa02 temporary adaption to new locale code
haftmann
parents: 29029
diff changeset
   598
    val mergeexpr = Locale.Merge suplocales;
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   599
    val constrain = Element.Constrains ((map o apsnd o map_atyps)
26167
ccc9007a7164 proper merge of base sorts
haftmann
parents: 25999
diff changeset
   600
      (K (TFree (Name.aT, base_sort))) supparams);
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   601
    fun fork_syn (Element.Fixes xs) =
29006
abe0f11cfa4e Name.name_of -> Binding.base_name
haftmann
parents: 28965
diff changeset
   602
          fold_map (fn (c, ty, syn) => cons (Binding.base_name c, syn) #> pair (c, ty, NoSyn)) xs
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   603
          #>> Element.Fixes
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   604
      | fork_syn x = pair x;
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   605
    fun fork_syntax elems =
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   606
      let
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   607
        val (elems', global_syntax) = fold_map fork_syn elems [];
26247
b6608fbeaae1 some theorems named explicitly
haftmann
parents: 26238
diff changeset
   608
      in (constrain :: elems', global_syntax) end;
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   609
    val (elems, global_syntax) =
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   610
      ProofContext.init thy
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   611
      |> Locale.cert_expr supexpr [constrain]
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   612
      |> snd
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   613
      |> init_ctxt sups base_sort
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   614
      |> process_expr Locale.empty raw_elems
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   615
      |> fst
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   616
      |> fork_syntax
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   617
  in (((sups, supparams), (supsort, base_sort, mergeexpr)), (elems, global_syntax)) end;
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   618
26247
b6608fbeaae1 some theorems named explicitly
haftmann
parents: 26238
diff changeset
   619
val read_class_spec = gen_class_spec Sign.intern_class Locale.read_expr;
b6608fbeaae1 some theorems named explicitly
haftmann
parents: 26238
diff changeset
   620
val check_class_spec = gen_class_spec (K I) Locale.cert_expr;
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   621
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   622
fun add_consts bname class base_sort sups supparams global_syntax thy =
24968
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   623
  let
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   624
    val supconsts = map fst supparams
26518
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   625
      |> AList.make (snd o the o AList.lookup (op =) (these_params thy sups))
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   626
      |> (map o apsnd o apsnd o map_atyps o K o TFree) (Name.aT, [class]);
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   627
    val all_params = map fst (Locale.parameters_of thy class);
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   628
    val raw_params = (snd o chop (length supparams)) all_params;
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   629
    fun add_const (v, raw_ty) thy =
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   630
      let
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28941
diff changeset
   631
        val c = Sign.full_bname thy v;
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   632
        val ty = map_atyps (K (TFree (Name.aT, base_sort))) raw_ty;
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   633
        val ty0 = Type.strip_sorts ty;
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   634
        val ty' = map_atyps (K (TFree (Name.aT, [class]))) ty0;
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   635
        val syn = (the_default NoSyn o AList.lookup (op =) global_syntax) v;
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   636
      in
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   637
        thy
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28941
diff changeset
   638
        |> Sign.declare_const [] ((Binding.name v, ty0), syn)
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   639
        |> snd
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   640
        |> pair ((v, ty), (c, ty'))
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   641
      end;
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   642
  in
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   643
    thy
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   644
    |> Sign.add_path (Logic.const_of_class bname)
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   645
    |> fold_map add_const raw_params
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   646
    ||> Sign.restore_naming thy
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   647
    |-> (fn params => pair (supconsts @ (map o apfst) fst params, params))
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   648
  end;
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   649
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   650
fun adjungate_axclass bname class base_sort sups supsort supparams global_syntax thy =
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   651
  let
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   652
    fun globalize param_map = map_aterms
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   653
      (fn Free (v, ty) => Const ((fst o the o AList.lookup (op =) param_map) v, ty)
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   654
        | t => t);
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   655
    val raw_pred = Locale.intros thy class
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   656
      |> fst
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   657
      |> map (Logic.unvarify o Logic.strip_imp_concl o Thm.prop_of);
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   658
    fun get_axiom thy = case (#axioms o AxClass.get_info thy) class
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   659
     of [] => NONE
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   660
      | [thm] => SOME thm;
24968
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   661
  in
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   662
    thy
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   663
    |> add_consts bname class base_sort sups supparams global_syntax
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   664
    |-> (fn (param_map, params) => AxClass.define_class (bname, supsort)
26518
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   665
          (map (fst o snd) params)
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28941
diff changeset
   666
          [((Binding.name (bname ^ "_" ^ AxClass.axiomsN), []), map (globalize param_map) raw_pred)]
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   667
    #> snd
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   668
    #> `get_axiom
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   669
    #-> (fn assm_axiom => fold (Sign.add_const_constraint o apsnd SOME o snd) params
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   670
    #> pair (map (Const o snd) param_map, param_map, params, assm_axiom)))
24968
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   671
  end;
f9bafc868847 replaced Sign.add_consts_authentic by Sign.declare_const;
wenzelm
parents: 24949
diff changeset
   672
26518
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   673
fun gen_class prep_spec bname raw_supclasses raw_elems thy =
24748
ee0a0eb6b738 proper syntax during class specification
haftmann
parents: 24731
diff changeset
   674
  let
28965
1de908189869 cleaned up binding module and related code
haftmann
parents: 28941
diff changeset
   675
    val class = Sign.full_bname thy bname;
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   676
    val (((sups, supparams), (supsort, base_sort, mergeexpr)), (elems, global_syntax)) =
26247
b6608fbeaae1 some theorems named explicitly
haftmann
parents: 26238
diff changeset
   677
      prep_spec thy raw_supclasses raw_elems;
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   678
    val supconsts = map (apsnd fst o snd) (these_params thy sups);
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   679
  in
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   680
    thy
28822
7ca11ecbc4fb explicit name morphism function for locale interpretation
haftmann
parents: 28739
diff changeset
   681
    |> Locale.add_locale "" bname mergeexpr elems
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   682
    |> snd
25311
aa750e3a581c Class.init now similiar to Locale.init
haftmann
parents: 25268
diff changeset
   683
    |> ProofContext.theory_of
26518
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   684
    |> adjungate_axclass bname class base_sort sups supsort supparams global_syntax
28715
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   685
    |-> (fn (inst, param_map, params, assm_axiom) =>
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   686
        `(fn thy => calculate_axiom thy sups base_sort assm_axiom param_map class)
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   687
    #-> (fn axiom =>
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   688
        prove_class_interpretation class inst
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   689
          (supconsts @ map (pair class o fst o snd) params) (the_list axiom) []
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   690
    #-> (fn morphism =>
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   691
        `(fn thy => activate_class_morphism thy class inst morphism)
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   692
    #-> (fn phi =>
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   693
        `(fn thy => calculate_rules thy phi sups base_sort param_map axiom class)
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   694
    #-> (fn (assm_intro, of_class) =>
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   695
        add_class_data ((class, sups), (params, base_sort, inst,
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   696
          morphism, axiom, assm_intro, of_class))
238f9966c80e class morphism stemming from locale interpretation
haftmann
parents: 28674
diff changeset
   697
    #> fold (note_assm_intro class) (the_list assm_intro))))))
25268
58146cb7bf44 more precise treatment of prove_subclass
haftmann
parents: 25239
diff changeset
   698
    |> init class
25038
522abf8a5f87 canonical interpretation interface
haftmann
parents: 25024
diff changeset
   699
    |> pair class
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   700
  end;
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   701
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   702
in
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   703
26518
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   704
val class_cmd = gen_class read_class_spec;
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   705
val class = gen_class check_class_spec;
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   706
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   707
end; (*local*)
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   708
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   709
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   710
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   711
(** instantiation target **)
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   712
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   713
(* bookkeeping *)
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   714
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   715
datatype instantiation = Instantiation of {
25864
11f531354852 explicit type variables for instantiation
haftmann
parents: 25829
diff changeset
   716
  arities: string list * (string * sort) list * sort,
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   717
  params: ((string * string) * (string * typ)) list
25603
4b7a58fc168c dropped Class.prep_spec
haftmann
parents: 25597
diff changeset
   718
    (*(instantiation parameter, type constructor), (local instantiation parameter, typ)*)
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   719
}
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   720
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   721
structure Instantiation = ProofDataFun
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   722
(
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   723
  type T = instantiation
25536
01753a944433 improved
haftmann
parents: 25518
diff changeset
   724
  fun init _ = Instantiation { arities = ([], [], []), params = [] };
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   725
);
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   726
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   727
fun mk_instantiation (arities, params) =
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   728
  Instantiation { arities = arities, params = params };
25514
4b508bb31a6c first working version of instance target
haftmann
parents: 25502
diff changeset
   729
fun get_instantiation lthy = case Instantiation.get (LocalTheory.target_of lthy)
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   730
 of Instantiation data => data;
25514
4b508bb31a6c first working version of instance target
haftmann
parents: 25502
diff changeset
   731
fun map_instantiation f = (LocalTheory.target o Instantiation.map)
4b508bb31a6c first working version of instance target
haftmann
parents: 25502
diff changeset
   732
  (fn Instantiation { arities, params } => mk_instantiation (f (arities, params)));
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   733
25514
4b508bb31a6c first working version of instance target
haftmann
parents: 25502
diff changeset
   734
fun the_instantiation lthy = case get_instantiation lthy
25536
01753a944433 improved
haftmann
parents: 25518
diff changeset
   735
 of { arities = ([], [], []), ... } => error "No instantiation target"
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   736
  | data => data;
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   737
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   738
val instantiation_params = #params o get_instantiation;
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   739
25514
4b508bb31a6c first working version of instance target
haftmann
parents: 25502
diff changeset
   740
fun instantiation_param lthy v = instantiation_params lthy
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   741
  |> find_first (fn (_, (v', _)) => v = v')
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   742
  |> Option.map (fst o fst);
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   743
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   744
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   745
(* syntax *)
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   746
26238
c30bb8182da2 generic improvable syntax for targets
haftmann
parents: 26167
diff changeset
   747
fun synchronize_inst_syntax ctxt =
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   748
  let
26259
d30f4a509361 better improvement in instantiation target
haftmann
parents: 26247
diff changeset
   749
    val Instantiation { arities = (_, _, sort), params = params } = Instantiation.get ctxt;
26238
c30bb8182da2 generic improvable syntax for targets
haftmann
parents: 26167
diff changeset
   750
    val thy = ProofContext.theory_of ctxt;
c30bb8182da2 generic improvable syntax for targets
haftmann
parents: 26167
diff changeset
   751
    fun subst (c, ty) = case AxClass.inst_tyco_of thy (c, ty)
c30bb8182da2 generic improvable syntax for targets
haftmann
parents: 26167
diff changeset
   752
         of SOME tyco => (case AList.lookup (op =) params (c, tyco)
c30bb8182da2 generic improvable syntax for targets
haftmann
parents: 26167
diff changeset
   753
             of SOME (v_ty as (_, ty)) => SOME (ty, Free v_ty)
c30bb8182da2 generic improvable syntax for targets
haftmann
parents: 26167
diff changeset
   754
              | NONE => NONE)
c30bb8182da2 generic improvable syntax for targets
haftmann
parents: 26167
diff changeset
   755
          | NONE => NONE;
c30bb8182da2 generic improvable syntax for targets
haftmann
parents: 26167
diff changeset
   756
    val unchecks =
c30bb8182da2 generic improvable syntax for targets
haftmann
parents: 26167
diff changeset
   757
      map (fn ((c, _), v_ty as (_, ty)) => (Free v_ty, Const (c, ty))) params;
c30bb8182da2 generic improvable syntax for targets
haftmann
parents: 26167
diff changeset
   758
  in
c30bb8182da2 generic improvable syntax for targets
haftmann
parents: 26167
diff changeset
   759
    ctxt
26259
d30f4a509361 better improvement in instantiation target
haftmann
parents: 26247
diff changeset
   760
    |> Overloading.map_improvable_syntax
26730
bbb5e6904d78 proper abbreviations in class
haftmann
parents: 26642
diff changeset
   761
         (fn (((primary_constraints, _), (((improve, _), _), _)), _) =>
bbb5e6904d78 proper abbreviations in class
haftmann
parents: 26642
diff changeset
   762
            (((primary_constraints, []), (((improve, subst), false), unchecks)), false))
26238
c30bb8182da2 generic improvable syntax for targets
haftmann
parents: 26167
diff changeset
   763
  end;
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   764
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   765
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   766
(* target *)
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   767
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   768
val sanatize_name = (*FIXME*)
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   769
  let
25574
016f677ad7b8 declaration of instance parameter names
haftmann
parents: 25536
diff changeset
   770
    fun is_valid s = Symbol.is_ascii_letter s orelse Symbol.is_ascii_digit s
016f677ad7b8 declaration of instance parameter names
haftmann
parents: 25536
diff changeset
   771
      orelse s = "'" orelse s = "_";
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   772
    val is_junk = not o is_valid andf Symbol.is_regular;
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   773
    val junk = Scan.many is_junk;
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   774
    val scan_valids = Symbol.scanner "Malformed input"
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   775
      ((junk |--
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   776
        (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
   777
        --| junk))
25999
f8bcd311d501 added ::: / @@@ scanner combinators;
wenzelm
parents: 25864
diff changeset
   778
      ::: Scan.repeat ((Scan.many1 is_valid >> implode) --| junk));
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   779
  in
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   780
    explode #> scan_valids #> implode
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   781
  end;
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   782
26259
d30f4a509361 better improvement in instantiation target
haftmann
parents: 26247
diff changeset
   783
fun type_name "*" = "prod"
d30f4a509361 better improvement in instantiation target
haftmann
parents: 26247
diff changeset
   784
  | type_name "+" = "sum"
d30f4a509361 better improvement in instantiation target
haftmann
parents: 26247
diff changeset
   785
  | type_name s = sanatize_name (NameSpace.base s); (*FIXME*)
d30f4a509361 better improvement in instantiation target
haftmann
parents: 26247
diff changeset
   786
26518
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   787
fun resort_terms pp algebra consts constraints ts =
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   788
  let
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   789
    fun matchings (Const (c_ty as (c, _))) = (case constraints c
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   790
         of NONE => I
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   791
          | SOME sorts => fold2 (curry (Sorts.meet_sort algebra))
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   792
              (Consts.typargs consts c_ty) sorts)
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   793
      | matchings _ = I
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   794
    val tvartab = (fold o fold_aterms) matchings ts Vartab.empty
26642
454d11701fa4 Sorts.class_error: produce message only (formerly msg_class_error);
wenzelm
parents: 26628
diff changeset
   795
      handle Sorts.CLASS_ERROR e => error (Sorts.class_error pp e);
27089
480f19078b65 fixed wrong treatment of type variables in instantiation target
haftmann
parents: 26995
diff changeset
   796
    val inst = map_type_tvar
480f19078b65 fixed wrong treatment of type variables in instantiation target
haftmann
parents: 26995
diff changeset
   797
      (fn (vi, sort) => TVar (vi, the_default sort (Vartab.lookup tvartab vi)));
26518
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   798
  in if Vartab.is_empty tvartab then NONE else SOME ((map o map_types) inst ts) end;
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   799
25864
11f531354852 explicit type variables for instantiation
haftmann
parents: 25829
diff changeset
   800
fun init_instantiation (tycos, vs, sort) thy =
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   801
  let
25536
01753a944433 improved
haftmann
parents: 25518
diff changeset
   802
    val _ = if null tycos then error "At least one arity must be given" else ();
26518
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   803
    val params = these_params thy sort;
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   804
    fun get_param tyco (param, (_, (c, ty))) = if can (AxClass.param_of_inst thy) (c, tyco)
25603
4b7a58fc168c dropped Class.prep_spec
haftmann
parents: 25597
diff changeset
   805
      then NONE else SOME ((c, tyco),
25864
11f531354852 explicit type variables for instantiation
haftmann
parents: 25829
diff changeset
   806
        (param ^ "_" ^ type_name tyco, map_atyps (K (Type (tyco, map TFree vs))) ty));
26518
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   807
    val inst_params = map_product get_param tycos params |> map_filter I;
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   808
    val primary_constraints = map (apsnd
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   809
      (map_atyps (K (TVar ((Name.aT, 0), [])))) o snd o snd) params;
26939
1035c89b4c02 moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents: 26761
diff changeset
   810
    val pp = Syntax.pp_global thy;
26518
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   811
    val algebra = Sign.classes_of thy
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   812
      |> fold (fn tyco => Sorts.add_arities pp
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   813
            (tyco, map (fn class => (class, map snd vs)) sort)) tycos;
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   814
    val consts = Sign.consts_of thy;
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   815
    val improve_constraints = AList.lookup (op =)
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   816
      (map (fn (_, (class, (c, _))) => (c, [[class]])) params);
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   817
    fun resort_check ts ctxt = case resort_terms pp algebra consts improve_constraints ts
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   818
     of NONE => NONE
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   819
      | SOME ts' => SOME (ts', ctxt);
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   820
    fun improve (c, ty) = case AxClass.inst_tyco_of thy (c, ty)
26329
3e58e4c67a2a instantiation less liberal with dangling constraints
haftmann
parents: 26259
diff changeset
   821
     of SOME tyco => (case AList.lookup (op =) inst_params (c, tyco)
26518
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   822
         of SOME (_, ty') => if Type.raw_instance (ty', ty) then SOME (ty, ty') else NONE
26259
d30f4a509361 better improvement in instantiation target
haftmann
parents: 26247
diff changeset
   823
          | NONE => NONE)
d30f4a509361 better improvement in instantiation target
haftmann
parents: 26247
diff changeset
   824
      | NONE => NONE;
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   825
  in
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   826
    thy
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   827
    |> ProofContext.init
26329
3e58e4c67a2a instantiation less liberal with dangling constraints
haftmann
parents: 26259
diff changeset
   828
    |> Instantiation.put (mk_instantiation ((tycos, vs, sort), inst_params))
27281
b457537e789a Variable.declare_typ;
wenzelm
parents: 27089
diff changeset
   829
    |> fold (Variable.declare_typ o TFree) vs
26329
3e58e4c67a2a instantiation less liberal with dangling constraints
haftmann
parents: 26259
diff changeset
   830
    |> fold (Variable.declare_names o Free o snd) inst_params
26259
d30f4a509361 better improvement in instantiation target
haftmann
parents: 26247
diff changeset
   831
    |> (Overloading.map_improvable_syntax o apfst)
26329
3e58e4c67a2a instantiation less liberal with dangling constraints
haftmann
parents: 26259
diff changeset
   832
         (fn ((_, _), ((_, subst), unchecks)) =>
26730
bbb5e6904d78 proper abbreviations in class
haftmann
parents: 26642
diff changeset
   833
            ((primary_constraints, []), (((improve, K NONE), false), [])))
26259
d30f4a509361 better improvement in instantiation target
haftmann
parents: 26247
diff changeset
   834
    |> Overloading.add_improvable_syntax
26518
3db6a46d8460 improved improvements for instantiaton
haftmann
parents: 26470
diff changeset
   835
    |> Context.proof_map (Syntax.add_term_check 0 "resorting" resort_check)
26238
c30bb8182da2 generic improvable syntax for targets
haftmann
parents: 26167
diff changeset
   836
    |> synchronize_inst_syntax
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   837
  end;
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   838
26238
c30bb8182da2 generic improvable syntax for targets
haftmann
parents: 26167
diff changeset
   839
fun confirm_declaration c = (map_instantiation o apsnd)
c30bb8182da2 generic improvable syntax for targets
haftmann
parents: 26167
diff changeset
   840
  (filter_out (fn (_, (c', _)) => c' = c))
c30bb8182da2 generic improvable syntax for targets
haftmann
parents: 26167
diff changeset
   841
  #> LocalTheory.target synchronize_inst_syntax
c30bb8182da2 generic improvable syntax for targets
haftmann
parents: 26167
diff changeset
   842
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   843
fun gen_instantiation_instance do_proof after_qed lthy =
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   844
  let
25864
11f531354852 explicit type variables for instantiation
haftmann
parents: 25829
diff changeset
   845
    val (tycos, vs, sort) = (#arities o the_instantiation) lthy;
11f531354852 explicit type variables for instantiation
haftmann
parents: 25829
diff changeset
   846
    val arities_proof = maps (fn tyco => Logic.mk_arities (tyco, map snd vs, sort)) tycos;
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   847
    fun after_qed' results =
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   848
      LocalTheory.theory (fold (AxClass.add_arity o Thm.varifyT) results)
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   849
      #> after_qed;
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   850
  in
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   851
    lthy
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   852
    |> do_proof after_qed' arities_proof
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   853
  end;
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   854
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   855
val instantiation_instance = gen_instantiation_instance (fn after_qed => fn ts =>
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   856
  Proof.theorem_i NONE (after_qed o map the_single) (map (fn t => [(t, [])]) ts));
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   857
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   858
fun prove_instantiation_instance tac = gen_instantiation_instance (fn after_qed =>
25502
9200b36280c0 instance command as rudimentary class target
haftmann
parents: 25485
diff changeset
   859
  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
   860
    (fn {context, ...} => tac context)) ts) lthy) I;
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   861
28666
d2dbfe3a0284 prove_instantiation_exit combinators
haftmann
parents: 28259
diff changeset
   862
fun prove_instantiation_exit tac = prove_instantiation_instance tac
d2dbfe3a0284 prove_instantiation_exit combinators
haftmann
parents: 28259
diff changeset
   863
  #> LocalTheory.exit_global;
d2dbfe3a0284 prove_instantiation_exit combinators
haftmann
parents: 28259
diff changeset
   864
d2dbfe3a0284 prove_instantiation_exit combinators
haftmann
parents: 28259
diff changeset
   865
fun prove_instantiation_exit_result f tac x lthy =
d2dbfe3a0284 prove_instantiation_exit combinators
haftmann
parents: 28259
diff changeset
   866
  let
d2dbfe3a0284 prove_instantiation_exit combinators
haftmann
parents: 28259
diff changeset
   867
    val phi = ProofContext.export_morphism lthy
d2dbfe3a0284 prove_instantiation_exit combinators
haftmann
parents: 28259
diff changeset
   868
      (ProofContext.init (ProofContext.theory_of lthy));
d2dbfe3a0284 prove_instantiation_exit combinators
haftmann
parents: 28259
diff changeset
   869
    val y = f phi x;
d2dbfe3a0284 prove_instantiation_exit combinators
haftmann
parents: 28259
diff changeset
   870
  in
d2dbfe3a0284 prove_instantiation_exit combinators
haftmann
parents: 28259
diff changeset
   871
    lthy
d2dbfe3a0284 prove_instantiation_exit combinators
haftmann
parents: 28259
diff changeset
   872
    |> prove_instantiation_exit (fn ctxt => tac ctxt y)
d2dbfe3a0284 prove_instantiation_exit combinators
haftmann
parents: 28259
diff changeset
   873
    |> pair y
d2dbfe3a0284 prove_instantiation_exit combinators
haftmann
parents: 28259
diff changeset
   874
  end;
d2dbfe3a0284 prove_instantiation_exit combinators
haftmann
parents: 28259
diff changeset
   875
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   876
fun conclude_instantiation lthy =
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   877
  let
25485
33840a854e63 tuned interfaces of class module
haftmann
parents: 25462
diff changeset
   878
    val { arities, params } = the_instantiation lthy;
25864
11f531354852 explicit type variables for instantiation
haftmann
parents: 25829
diff changeset
   879
    val (tycos, vs, sort) = arities;
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   880
    val thy = ProofContext.theory_of lthy;
25597
34860182b250 moved instance parameter management from class.ML to axclass.ML
haftmann
parents: 25574
diff changeset
   881
    val _ = map (fn tyco => if Sign.of_sort thy
25864
11f531354852 explicit type variables for instantiation
haftmann
parents: 25829
diff changeset
   882
        (Type (tyco, map TFree vs), sort)
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   883
      then () else error ("Missing instance proof for type " ^ quote (Sign.extern_type thy tyco)))
25597
34860182b250 moved instance parameter management from class.ML to axclass.ML
haftmann
parents: 25574
diff changeset
   884
        tycos;
34860182b250 moved instance parameter management from class.ML to axclass.ML
haftmann
parents: 25574
diff changeset
   885
  in lthy end;
25462
dad0291cb76a rudimentary instantiation target
haftmann
parents: 25368
diff changeset
   886
25603
4b7a58fc168c dropped Class.prep_spec
haftmann
parents: 25597
diff changeset
   887
fun pretty_instantiation lthy =
4b7a58fc168c dropped Class.prep_spec
haftmann
parents: 25597
diff changeset
   888
  let
4b7a58fc168c dropped Class.prep_spec
haftmann
parents: 25597
diff changeset
   889
    val { arities, params } = the_instantiation lthy;
25864
11f531354852 explicit type variables for instantiation
haftmann
parents: 25829
diff changeset
   890
    val (tycos, vs, sort) = arities;
25603
4b7a58fc168c dropped Class.prep_spec
haftmann
parents: 25597
diff changeset
   891
    val thy = ProofContext.theory_of lthy;
25864
11f531354852 explicit type variables for instantiation
haftmann
parents: 25829
diff changeset
   892
    fun pr_arity tyco = Syntax.pretty_arity lthy (tyco, map snd vs, sort);
25603
4b7a58fc168c dropped Class.prep_spec
haftmann
parents: 25597
diff changeset
   893
    fun pr_param ((c, _), (v, ty)) =
25864
11f531354852 explicit type variables for instantiation
haftmann
parents: 25829
diff changeset
   894
      (Pretty.block o Pretty.breaks) [Pretty.str v, Pretty.str "==",
26939
1035c89b4c02 moved global pretty/string_of functions from Sign to Syntax;
wenzelm
parents: 26761
diff changeset
   895
        (Pretty.str o Sign.extern_const thy) c, Pretty.str "::", Syntax.pretty_typ_global thy ty];
25603
4b7a58fc168c dropped Class.prep_spec
haftmann
parents: 25597
diff changeset
   896
  in
4b7a58fc168c dropped Class.prep_spec
haftmann
parents: 25597
diff changeset
   897
    (Pretty.block o Pretty.fbreaks)
4b7a58fc168c dropped Class.prep_spec
haftmann
parents: 25597
diff changeset
   898
      (Pretty.str "instantiation" :: map pr_arity tycos @ map pr_param params)
4b7a58fc168c dropped Class.prep_spec
haftmann
parents: 25597
diff changeset
   899
  end;
4b7a58fc168c dropped Class.prep_spec
haftmann
parents: 25597
diff changeset
   900
24218
fbf1646b267c ClassPackage renamed to Class
haftmann
parents:
diff changeset
   901
end;
25683
d9fefc4859be maior tuning
haftmann
parents: 25668
diff changeset
   902